summaryrefslogtreecommitdiff
path: root/pretyping
diff options
context:
space:
mode:
authorGravatar Samuel Mimram <smimram@debian.org>2006-04-28 14:59:16 +0000
committerGravatar Samuel Mimram <smimram@debian.org>2006-04-28 14:59:16 +0000
commit3ef7797ef6fc605dfafb32523261fe1b023aeecb (patch)
treead89c6bb57ceee608fcba2bb3435b74e0f57919e /pretyping
parent018ee3b0c2be79eb81b1f65c3f3fa142d24129c8 (diff)
Imported Upstream version 8.0pl3+8.1alphaupstream/8.0pl3+8.1alpha
Diffstat (limited to 'pretyping')
-rw-r--r--pretyping/cases.ml490
-rw-r--r--pretyping/cases.mli36
-rw-r--r--pretyping/cbv.ml12
-rw-r--r--pretyping/cbv.mli2
-rw-r--r--[-rwxr-xr-x]pretyping/classops.ml188
-rw-r--r--pretyping/classops.mli22
-rw-r--r--pretyping/clenv.ml435
-rw-r--r--pretyping/clenv.mli117
-rw-r--r--pretyping/coercion.ml425
-rw-r--r--pretyping/coercion.mli58
-rw-r--r--pretyping/detyping.ml429
-rw-r--r--pretyping/detyping.mli25
-rw-r--r--pretyping/evarconv.ml415
-rw-r--r--pretyping/evarconv.mli19
-rw-r--r--pretyping/evarutil.ml721
-rw-r--r--pretyping/evarutil.mli143
-rw-r--r--pretyping/evd.ml522
-rw-r--r--pretyping/evd.mli123
-rw-r--r--pretyping/indrec.ml249
-rw-r--r--pretyping/indrec.mli27
-rw-r--r--pretyping/inductiveops.ml145
-rw-r--r--pretyping/inductiveops.mli23
-rw-r--r--pretyping/instantiate.ml68
-rw-r--r--pretyping/instantiate.mli25
-rw-r--r--pretyping/matching.ml35
-rw-r--r--pretyping/matching.mli11
-rw-r--r--pretyping/pattern.ml160
-rw-r--r--pretyping/pattern.mli21
-rw-r--r--pretyping/pretype_errors.ml27
-rw-r--r--pretyping/pretype_errors.mli24
-rw-r--r--pretyping/pretyping.ml1617
-rw-r--r--pretyping/pretyping.mli142
-rw-r--r--pretyping/rawterm.ml197
-rw-r--r--pretyping/rawterm.mli43
-rw-r--r--[-rwxr-xr-x]pretyping/recordops.ml230
-rwxr-xr-xpretyping/recordops.mli38
-rw-r--r--pretyping/reductionops.ml214
-rw-r--r--pretyping/reductionops.mli17
-rw-r--r--pretyping/retyping.ml66
-rw-r--r--pretyping/retyping.mli12
-rw-r--r--pretyping/tacred.ml263
-rw-r--r--pretyping/tacred.mli24
-rw-r--r--pretyping/termops.ml178
-rw-r--r--pretyping/termops.mli29
-rw-r--r--pretyping/typing.ml186
-rw-r--r--pretyping/typing.mli25
-rw-r--r--pretyping/unification.ml471
-rw-r--r--pretyping/unification.mli33
48 files changed, 5441 insertions, 3341 deletions
diff --git a/pretyping/cases.ml b/pretyping/cases.ml
index 4aff508f..a32aaf45 100644
--- a/pretyping/cases.ml
+++ b/pretyping/cases.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: cases.ml,v 1.111.2.5 2005/04/29 16:31:03 herbelin Exp $ *)
+(* $Id: cases.ml 8693 2006-04-10 12:05:05Z msozeau $ *)
open Util
open Names
@@ -33,6 +33,7 @@ type pattern_matching_error =
| BadPattern of constructor * constr
| BadConstructor of constructor * inductive
| WrongNumargConstructor of constructor * int
+ | WrongNumargInductive of inductive * int
| WrongPredicateArity of constr * constr * constr
| NeedsInversion of constr * constr
| UnusedClause of cases_pattern list
@@ -50,8 +51,11 @@ let error_bad_pattern_loc loc cstr ind =
let error_bad_constructor_loc loc cstr ind =
raise_pattern_matching_error (loc, Global.env(), BadConstructor (cstr,ind))
-let error_wrong_numarg_constructor_loc loc c n =
- raise_pattern_matching_error (loc, Global.env(), WrongNumargConstructor (c,n))
+let error_wrong_numarg_constructor_loc loc env c n =
+ raise_pattern_matching_error (loc, env, WrongNumargConstructor(c,n))
+
+let error_wrong_numarg_inductive_loc loc env c n =
+ raise_pattern_matching_error (loc, env, WrongNumargInductive(c,n))
let error_wrong_predicate_arity_loc loc env c n1 n2 =
raise_pattern_matching_error (loc, env, WrongPredicateArity (c,n1,n2))
@@ -59,107 +63,18 @@ let error_wrong_predicate_arity_loc loc env c n1 n2 =
let error_needs_inversion env x t =
raise (PatternMatchingError (env, NeedsInversion (x,t)))
-(*********************************************************************)
-(* A) Typing old cases *)
-(* This was previously in Indrec but creates existential holes *)
-
-let mkExistential isevars env loc = new_isevar isevars env loc (new_Type ())
-
-let norec_branch_scheme env isevars cstr =
- let rec crec env = function
- | d::rea -> mkProd_or_LetIn d (crec (push_rel d env) rea)
- | [] -> mkExistential isevars env (dummy_loc, InternalHole) in
- crec env (List.rev cstr.cs_args)
-
-let rec_branch_scheme env isevars (sp,j) recargs cstr =
- let rec crec env (args,recargs) =
- match args, recargs with
- | (name,None,c as d)::rea,(ra::reca) ->
- let d =
- match dest_recarg ra with
- | Mrec k when k=j ->
- let t = mkExistential isevars env (dummy_loc, InternalHole)
- in
- mkArrow t
- (crec (push_rel (Anonymous,None,t) env)
- (List.rev (lift_rel_context 1 (List.rev rea)),reca))
- | _ -> crec (push_rel d env) (rea,reca) in
- mkProd (name, c, d)
-
- | (name,Some b,c as d)::rea, reca ->
- mkLetIn (name,b, c,crec (push_rel d env) (rea,reca))
- | [],[] -> mkExistential isevars env (dummy_loc, InternalHole)
- | _ -> anomaly "rec_branch_scheme"
- in
- crec env (List.rev cstr.cs_args,recargs)
-
-let branch_scheme env isevars isrec indf =
- let (ind,params) = dest_ind_family indf in
- let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let cstrs = get_constructors env indf in
- if isrec then
- array_map2
- (rec_branch_scheme env isevars ind)
- (dest_subterms mip.mind_recargs) cstrs
- else
- Array.map (norec_branch_scheme env isevars) cstrs
-
-(******************************************************)
-(* B) Building ML like case expressions without types *)
-
-let concl_n env sigma =
- let rec decrec m c = if m = 0 then (nf_evar sigma c) else
- match kind_of_term (whd_betadeltaiota env sigma c) with
- | Prod (n,_,c_0) -> decrec (m-1) c_0
- | _ -> failwith "Typing.concl_n"
- in
- decrec
-
-let count_rec_arg j =
- let rec crec i = function
- | [] -> i
- | ra::l ->
- (match dest_recarg ra with
- Mrec k -> crec (if k=j then (i+1) else i) l
- | _ -> crec i l)
- in
- crec 0
-
-(* if arity of mispec is (p_bar:P_bar)(a_bar:A_bar)s where p_bar are the
- * K parameters. Then then build_notdep builds the predicate
- * [a_bar:A'_bar](lift k pred)
- * where A'_bar = A_bar[p_bar <- globargs] *)
-
-let build_dep_pred env sigma indf pred =
- let arsign,_ = get_arity env indf in
- let psign = (Anonymous,None,build_dependent_inductive env indf)::arsign in
- let nar = List.length psign in
- it_mkLambda_or_LetIn_name env (lift nar pred) psign
-
-type ml_case_error =
- | MlCaseAbsurd
- | MlCaseDependent
-
-exception NotInferable of ml_case_error
-
-
-let pred_case_ml env sigma isrec (IndType (indf,realargs)) (i,ft) =
- let pred =
- let (ind,params) = dest_ind_family indf in
- let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let recargs = dest_subterms mip.mind_recargs in
- if Array.length recargs = 0 then raise (NotInferable MlCaseAbsurd);
- let recargi = recargs.(i) in
- let j = snd ind in (* index of inductive *)
- let nbrec = if isrec then count_rec_arg j recargi else 0 in
- let nb_arg = List.length (recargs.(i)) + nbrec in
- let pred = Evarutil.refresh_universes (concl_n env sigma nb_arg ft) in
- if noccur_between 1 nb_arg pred then
- lift (-nb_arg) pred
- else
- raise (NotInferable MlCaseDependent)
- in
- build_dep_pred env sigma indf pred
+module type S = sig
+ val compile_cases :
+ loc ->
+ (type_constraint -> env -> rawconstr -> unsafe_judgment) *
+ Evd.evar_defs ref ->
+ type_constraint ->
+ env ->
+ rawconstr option *
+ (rawconstr * (name * (loc * inductive * name list) option)) list *
+ (loc * identifier list * cases_pattern list * rawconstr) list ->
+ unsafe_judgment
+end
(************************************************************************)
(* Pattern-matching compilation (Cases) *)
@@ -259,8 +174,8 @@ type tomatch_stack = tomatch_status list
(* The type [predicate_signature] types the terms to match and the rhs:
- - [PrLetIn (names,dep,pred)] types a pushed term ([Pushed]),
- if dep<>Anonymous, the term is dependent, let n=|names|, if
+ - [PrLetIn (names,dep,pred)] types a pushed term ([Pushed]),
+ if dep<>Anonymous, the term is dependent, let n=|names|, if
n<>0 then the type of the pushed term is necessarily an
inductive with n real arguments. Otherwise, it may be
non inductive, or inductive without real arguments, or inductive
@@ -376,7 +291,7 @@ let push_history_pattern n current cont =
*)
type pattern_matching_problem =
{ env : env;
- isevars : evar_defs;
+ isevars : Evd.evar_defs ref;
pred : predicate_signature option;
tomatch : tomatch_stack;
history : pattern_continuation;
@@ -388,64 +303,67 @@ type pattern_matching_problem =
* A few functions to infer the inductive type from the patterns instead of *
* checking that the patterns correspond to the ind. type of the *
* destructurated object. Allows type inference of examples like *
- * [n]Cases n of O => true | _ => false end *
+ * match n with O => true | _ => false end *
+ * match x in I with C => true | _ => false end *
*--------------------------------------------------------------------------*)
(* Computing the inductive type from the matrix of patterns *)
+(* We use the "in I" clause to coerce the terms to match and otherwise
+ use the constructor to know in which type is the matching problem
+
+ Note that insertion of coercions inside nested patterns is done
+ each time the matrix is expanded *)
+
let rec find_row_ind = function
[] -> None
| PatVar _ :: l -> find_row_ind l
| PatCstr(loc,c,_,_) :: _ -> Some (loc,c)
-exception NotCoercible
-
-let inh_coerce_to_ind isevars env tmloc ty tyi =
- let (mib,mip) = Inductive.lookup_mind_specif env tyi in
- let (ntys,_) = splay_prod env (evars_of isevars) mip.mind_nf_arity in
+let inductive_template isevars env tmloc ind =
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ let (ntys,_) = splay_prod env (Evd.evars_of !isevars) mip.mind_nf_arity in
let hole_source = match tmloc with
- | Some loc -> fun i -> (loc, TomatchTypeParameter (tyi,i))
- | None -> fun _ -> (dummy_loc, InternalHole) in
+ | Some loc -> fun i -> (loc, Evd.TomatchTypeParameter (ind,i))
+ | None -> fun _ -> (dummy_loc, Evd.InternalHole) in
let (evarl,_) =
List.fold_right
(fun (na,ty) (evl,n) ->
- (new_isevar isevars env (hole_source n) (substl evl ty))::evl,n+1)
+ (e_new_evar isevars env ~src:(hole_source n) (substl evl ty))::evl,n+1)
ntys ([],1) in
- let expected_typ = applist (mkInd tyi,evarl) in
+ applist (mkInd ind,List.rev evarl)
+
+let inh_coerce_to_ind isevars env ty tyi =
+ let expected_typ = inductive_template isevars env None tyi in
(* devrait être indifférent d'exiger leq ou pas puisque pour
un inductif cela doit être égal *)
- if the_conv_x_leq env isevars expected_typ ty then ty
- else raise NotCoercible
-
-(* We do the unification for all the rows that contain
- * constructor patterns. This is what we do at the higher level of patterns.
- * For nested patterns, we do this unif when we ``expand'' the matrix, and we
- * use the function above.
- *)
-
-let unify_tomatch_with_patterns isevars env tmloc typ = function
- | Some (cloc,(cstr,_ as c)) ->
- (let tyi = inductive_of_constructor c in
- try
- let indtyp = inh_coerce_to_ind isevars env tmloc typ tyi in
- IsInd (typ,find_rectype env (evars_of isevars) typ)
- with NotCoercible ->
- (* 2 cases : Not the right inductive or not an inductive at all *)
- try
- IsInd (typ,find_rectype env (evars_of isevars) typ)
- (* will try to coerce later in check_and_adjust_constructor.. *)
- with Not_found ->
- NotInd (None,typ))
- (* error will be detected in check_all_variables *)
- | None ->
- try IsInd (typ,find_rectype env (evars_of isevars) typ)
- with Not_found -> NotInd (None,typ)
-
-let coerce_row typing_fun isevars env cstropt tomatch =
- let j = typing_fun empty_tycon env tomatch in
- let typ = body_of_type j.uj_type in
- let loc = loc_of_rawconstr tomatch in
- let t = unify_tomatch_with_patterns isevars env (Some loc) typ cstropt in
+ let _ = e_cumul env isevars expected_typ ty in ()
+
+let unify_tomatch_with_patterns isevars env typ tm =
+ match find_row_ind tm with
+ | None -> NotInd (None,typ)
+ | Some (_,(ind,_)) ->
+ inh_coerce_to_ind isevars env typ ind;
+ try IsInd (typ,find_rectype env (Evd.evars_of !isevars) typ)
+ with Not_found -> NotInd (None,typ)
+
+let find_tomatch_tycon isevars env loc = function
+ (* Try first if some 'in I ...' is present and can be used as a constraint *)
+ | Some (_,ind,_),_
+ (* Otherwise try to get constraints from (the 1st) constructor in clauses *)
+ | None, Some (_,(ind,_)) ->
+ mk_tycon (inductive_template isevars env loc ind)
+ | None, None ->
+ empty_tycon
+
+let coerce_row typing_fun isevars env cstropt (tomatch,(_,indopt)) =
+ let loc = Some (loc_of_rawconstr tomatch) in
+ let tycon = find_tomatch_tycon isevars env loc (indopt,cstropt) in
+ let j = typing_fun tycon env tomatch in
+ let typ = nf_evar (Evd.evars_of !isevars) j.uj_type in
+ let t =
+ try IsInd (typ,find_rectype env (Evd.evars_of !isevars) typ)
+ with Not_found -> NotInd (None,typ) in
(j.uj_val,t)
let coerce_to_indtype typing_fun isevars env matx tomatchl =
@@ -458,11 +376,53 @@ let coerce_to_indtype typing_fun isevars env matx tomatchl =
(************************************************************************)
(* Utils *)
+let mkExistential env ?(src=(dummy_loc,Evd.InternalHole)) isevars =
+ e_new_evar isevars env ~src:src (new_Type ())
+
+let evd_comb2 f isevars x y =
+ let (evd',y) = f !isevars x y in
+ isevars := evd';
+ y
+
+
+module Cases_F(Coercion : Coercion.S) : S = struct
+
+let adjust_tomatch_to_pattern pb ((current,typ),deps) =
+ (* Ideally, we could find a common inductive type to which both the
+ term to match and the patterns coerce *)
+ (* In practice, we coerce the term to match if it is not already an
+ inductive type and it is not dependent; moreover, we use only
+ the first pattern type and forget about the others *)
+ let typ = match typ with IsInd (t,_) -> t | NotInd (_,t) -> t in
+ let typ =
+ try IsInd (typ,find_rectype pb.env (Evd.evars_of !(pb.isevars)) typ)
+ with Not_found -> NotInd (None,typ) in
+ let tomatch = ((current,typ),deps) in
+ match typ with
+ | NotInd (None,typ) ->
+ let tm1 = List.map (fun eqn -> List.hd eqn.patterns) pb.mat in
+ (match find_row_ind tm1 with
+ | None -> tomatch
+ | Some (_,(ind,_)) ->
+ let indt = inductive_template pb.isevars pb.env None ind in
+ let current =
+ if deps = [] & isEvar typ then
+ (* Don't insert coercions if dependent; only solve evars *)
+ let _ = e_cumul pb.env pb.isevars indt typ in
+ current
+ else
+ (evd_comb2 (Coercion.inh_conv_coerce_to dummy_loc pb.env)
+ pb.isevars (make_judge current typ) (mk_tycon_type indt)).uj_val in
+ let sigma = Evd.evars_of !(pb.isevars) in
+ let typ = IsInd (indt,find_rectype pb.env sigma indt) in
+ ((current,typ),deps))
+ | _ -> tomatch
+
(* extract some ind from [t], possibly coercing from constructors in [tm] *)
let to_mutind env isevars tm c t =
- match c with
- | Some body -> NotInd (c,t)
- | None -> unify_tomatch_with_patterns isevars env None t (find_row_ind tm)
+(* match c with
+ | Some body -> *) NotInd (c,t)
+(* | None -> unify_tomatch_with_patterns isevars env t tm*)
let type_of_tomatch = function
| IsInd (t,_) -> t
@@ -558,7 +518,8 @@ let check_and_adjust_constructor ind cstrs = function
let args' = adjust_local_defs loc (args, List.rev ci.cs_args)
in PatCstr (loc, cstr, args', alias)
with NotAdjustable ->
- error_wrong_numarg_constructor_loc loc cstr nb_args_constr
+ error_wrong_numarg_constructor_loc loc (Global.env())
+ cstr nb_args_constr
else
(* Try to insert a coercion *)
try
@@ -715,11 +676,11 @@ let lift_tomatch_stack n = liftn_tomatch_stack n 1
(* Some heuristics to get names for variables pushed in pb environment *)
(* Typical requirement:
- [Cases y of (S (S x)) => x | x => x end] should be compiled into
- [Cases y of O => y | (S n) => Cases n of O => y | (S x) => x end end]
+ [match y with (S (S x)) => x | x => x end] should be compiled into
+ [match y with O => y | (S n) => match n with O => y | (S x) => x end end]
- and [Cases y of (S (S n)) => n | n => n end] into
- [Cases y of O => y | (S n0) => Cases n0 of O => y | (S n) => n end end]
+ and [match y with (S (S n)) => n | n => n end] into
+ [match y with O => y | (S n0) => match n0 with O => y | (S n) => n end end]
i.e. user names should be preserved and created names should not
interfere with user names *)
@@ -782,7 +743,6 @@ let build_aliases_context env sigma names allpats pats =
let newallpats =
List.map2 (fun l1 l2 -> List.hd l2::l1) newallpats oldallpats in
let oldallpats = List.map List.tl oldallpats in
- let u = Retyping.get_type_of env sigma deppat in
let decl = (na,Some deppat,t) in
let a = (deppat,nondeppat,d,t) in
insert (push_rel decl env) (decl::sign1) ((na,a)::sign2) (n+1)
@@ -869,10 +829,10 @@ the following n+1 equations:
Some hints:
-- Clearly, if xij occurs in Ti, then, a "Cases z of (Ci xi1..xipi) => ..."
+- Clearly, if xij occurs in Ti, then, a "match z with (Ci xi1..xipi) => ..."
should be inserted somewhere in Ti.
-- If T is undefined, an easy solution is to insert a "Cases z of (Ci
+- If T is undefined, an easy solution is to insert a "match z with (Ci
xi1..xipi) => ..." in front of each Ti
- Otherwise, T1..Tn and T must be step by step unified, if some of them
@@ -902,7 +862,7 @@ let shift_operator k = function OpLambda _ | OpProd _ -> k+1 | _ -> k
let reloc_operator (k,n) = function OpRel p when p > k ->
let rec unify_clauses k pv =
- let pv'= Array.map (fun (n,sign,_,p) -> n,splay_constr (whd_betaiotaevar (push_rels (List.rev sign) env) (evars_of isevars)) p) pv in
+ let pv'= Array.map (fun (n,sign,_,p) -> n,splay_constr (whd_betaiotaevar (push_rels (List.rev sign) env) (Evd.evars_of isevars)) p) pv in
let n1,op1 = let (n1,(op1,args1)) = pv'.(0) in n1,op1 in
if Array.for_all (fun (ni,(opi,_)) -> eq_operator_lift k (n1,ni) (op1,opi)) pv'
then
@@ -918,28 +878,30 @@ let abstract_conclusion typ cs =
lam_it p sign
let infer_predicate loc env isevars typs cstrs indf =
- let (mis,_) = dest_ind_family indf in
(* Il faudra substituer les isevars a un certain moment *)
if Array.length cstrs = 0 then (* "TODO4-3" *)
error "Inference of annotation for empty inductive types not implemented"
else
(* Empiric normalization: p may depend in a irrelevant way on args of the*)
- (* cstr as in [c:{_:Alpha & Beta}] Cases c of (existS a b)=>(a,b) end *)
+ (* cstr as in [c:{_:Alpha & Beta}] match c with (existS a b)=>(a,b) end *)
let typs =
- Array.map (local_strong (whd_betaevar empty_env (evars_of isevars))) typs
+ Array.map (local_strong (whd_betaevar empty_env (Evd.evars_of !isevars))) typs
in
let eqns = array_map2 prepare_unif_pb typs cstrs in
(* First strategy: no dependencies at all *)
-(* let (cclargs,_,typn) = eqns.(mis_nconstr mis -1) in*)
+(*
+ let (mis,_) = dest_ind_family indf in
+ let (cclargs,_,typn) = eqns.(mis_nconstr mis -1) in
+*)
let (sign,_) = get_arity env indf in
let mtyp =
if array_exists is_Type typs then
(* Heuristic to avoid comparison between non-variables algebric univs*)
new_Type ()
else
- mkExistential isevars env (loc, CasesType)
+ mkExistential env ~src:(loc, Evd.CasesType) isevars
in
- if array_for_all (fun (_,_,typ) -> the_conv_x_leq env isevars typ mtyp) eqns
+ if array_for_all (fun (_,_,typ) -> e_cumul env isevars typ mtyp) eqns
then
(* Non dependent case -> turn it into a (dummy) dependent one *)
let sign = (Anonymous,None,build_dependent_inductive env indf)::sign in
@@ -980,7 +942,7 @@ let rec map_predicate f k = function
let rec noccurn_predicate k = function
| PrCcl ccl -> noccurn k ccl
| PrProd pred -> noccurn_predicate (k+1) pred
- | PrLetIn ((names,dep as tm),pred) ->
+ | PrLetIn ((names,dep),pred) ->
let k' = List.length names + (if dep<>Anonymous then 1 else 0) in
noccurn_predicate (k+k') pred
@@ -1066,8 +1028,8 @@ let abstract_predicate env sigma indf cur tms = function
args or not, we lift it to make room for [sign] *)
(* Even if not intrinsically dep, we move the predicate into a dep one *)
let sign,k =
- if names = [] & n <> 1 then
- (* Real args were not considered *)
+ if names = [] & n <> 1 then
+ (* Real args were not considered *)
(if dep<>Anonymous then
((let (_,c,t) = List.hd sign in (dep,c,t)::List.tl sign),n-1)
else
@@ -1106,7 +1068,7 @@ let expand_arg n alreadydep (na,t) deps (k,pred) =
(*****************************************************************************)
(* pred = [X:=realargs;x:=c]P types the following problem: *)
(* *)
-(* Gamma |- Cases Pushed(c:I(realargs)) rest of...end: pred *)
+(* Gamma |- match Pushed(c:I(realargs)) rest with...end: pred *)
(* *)
(* where the branch with constructor Ci:(x1:T1)...(xn:Tn)->I(realargsi) *)
(* is considered. Assume each Ti is some Ii(argsi). *)
@@ -1114,7 +1076,7 @@ let expand_arg n alreadydep (na,t) deps (k,pred) =
(* *)
(* pred' = [X1:=rargs1,x1:=x1']...[Xn:=rargsn,xn:=xn'](P[X:=realargsi;x:=e]) *)
(* *)
-(* s.t Gamma,x1'..xn' |- Cases Pushed(x1')..Pushed(xn') rest of...end: pred' *)
+(* s.t Gamma,x1'..xn' |- match Pushed(x1')..Pushed(xn') rest with..end :pred'*)
(* *)
(*****************************************************************************)
let specialize_predicate tomatchs deps cs = function
@@ -1141,7 +1103,7 @@ let find_predicate loc env isevars p typs cstrs current
(IndType (indf,realargs)) tms =
let (dep,pred) =
match p with
- | Some p -> abstract_predicate env (evars_of isevars) indf current tms p
+ | Some p -> abstract_predicate env (Evd.evars_of !isevars) indf current tms p
| None -> infer_predicate loc env isevars typs cstrs indf in
let typ = whd_beta (applist (pred, realargs)) in
if dep then
@@ -1191,7 +1153,7 @@ let group_equations pb mind current cstrs mat =
let rest = {rest with tag = lower_pattern_status rest.tag} in
brs.(i-1) <- (args, rest) :: brs.(i-1)
done
- | PatCstr (loc,((_,i) as cstr),args,_) as pat ->
+ | PatCstr (loc,((_,i)),args,_) ->
(* This is a regular clause *)
only_default := false;
brs.(i-1) <- (args,rest) :: brs.(i-1)) mat () in
@@ -1240,8 +1202,6 @@ let build_branch current deps pb eqns const_info =
else
DepAlias
in
- let partialci =
- applist (mkConstruct const_info.cs_cstr, const_info.cs_params) in
let history =
push_history_pattern const_info.cs_nargs
(AliasConstructor const_info.cs_cstr)
@@ -1324,10 +1284,8 @@ let rec compile pb =
| (Abstract d)::rest -> compile_generalization pb d rest
| [] -> build_leaf pb
-and match_current pb ((current,typ as ct),deps) =
- let tm1 = List.map (fun eqn -> List.hd eqn.patterns) pb.mat in
- let (_,c,t) = mkDeclTomatch Anonymous typ in
- let typ = to_mutind pb.env pb.isevars tm1 c t in
+and match_current pb tomatch =
+ let ((current,typ as ct),deps) = adjust_tomatch_to_pattern pb tomatch in
match typ with
| NotInd (_,typ) ->
check_all_variables typ pb.mat;
@@ -1336,10 +1294,10 @@ and match_current pb ((current,typ as ct),deps) =
let mind,_ = dest_ind_family indf in
let cstrs = get_constructors pb.env indf in
let eqns,onlydflt = group_equations pb mind current cstrs pb.mat in
- if (cstrs <> [||] or not (initial_history pb.history)) & onlydflt then
+ if (Array.length cstrs <> 0 or pb.mat <> []) & onlydflt then
compile (shift_problem ct pb)
else
- let constraints = Array.map (solve_constraints indt) cstrs in
+ let _constraints = Array.map (solve_constraints indt) cstrs in
(* We generalize over terms depending on current term to match *)
let pb = generalize_problem pb current deps in
@@ -1381,7 +1339,7 @@ and compile_generalization pb d rest =
and compile_alias pb (deppat,nondeppat,d,t) rest =
let history = simplify_history pb.history in
let sign, newenv, mat =
- insert_aliases pb.env (evars_of pb.isevars) (deppat,nondeppat,d,t) pb.mat in
+ insert_aliases pb.env (Evd.evars_of !(pb.isevars)) (deppat,nondeppat,d,t) pb.mat in
let n = List.length sign in
(* We had Gamma1; x:current; Gamma2 |- tomatch(x) and we rebind x to get *)
@@ -1436,7 +1394,7 @@ let rename_env subst env =
let is_dependent_indtype = function
| NotInd _ -> false
- | IsInd (_, IndType(_,realargs)) -> List.length realargs <> 0
+ | IsInd (_, IndType(_,realargs)) -> realargs <> []
let prepare_initial_alias_eqn isdep tomatchl eqn =
let (subst, pats) =
@@ -1444,7 +1402,7 @@ let prepare_initial_alias_eqn isdep tomatchl eqn =
(fun pat (tm,tmtyp) (subst, stripped_pats) ->
match alias_of_pat pat with
| Anonymous -> (subst, pat::stripped_pats)
- | Name idpat as na ->
+ | Name idpat ->
match kind_of_term tm with
| Rel n when not (is_dependent_indtype tmtyp) & not isdep
-> (n, idpat)::subst, (unalias_pat pat::stripped_pats)
@@ -1531,7 +1489,7 @@ let extract_predicate_conclusion isdep tomatchl pred =
if n=0 then (l,p) else
match kind_of_term p with
| Lambda (na,_,c) -> decomp_lam_force (n-1) (na::l) c
- | _ -> (* eta-expansion *)
+ | _ -> (* eta-expansion *)
let na = Name (id_of_string "x") in
decomp_lam_force (n-1) (na::l) (applist (lift 1 p, [mkRel 1])) in
let rec buildrec allnames p = function
@@ -1543,7 +1501,7 @@ let extract_predicate_conclusion isdep tomatchl pred =
(* adjust to a sign containing the NotInd's *)
if isdep then lift 1 p else p in
let names = if isdep then [Anonymous] else [] in
- buildrec (names::allnames) p ltm
+ buildrec (names::allnames) p ltm
| Some n ->
let n = if isdep then n+1 else n in
let names,p = decomp_lam_force n [] p in
@@ -1590,30 +1548,32 @@ let set_arity_signature dep n arsign tomatchl pred x =
decomp_block [] pred (tomatchl,arsign)
let prepare_predicate_from_tycon loc dep env isevars tomatchs c =
- let cook (n, l, signs) = function
+ let cook (n, l, env, signs) = function
| c,IsInd (_,IndType(indf,realargs)) ->
let indf' = lift_inductive_family n indf in
- let arsign = make_arity_signature env dep indf' in
+ let sign = make_arity_signature env dep indf' in
let p = List.length realargs in
if dep then
- (n + p + 1, c::(List.rev realargs)@l, arsign::signs)
+ (n + p + 1, c::(List.rev realargs)@l, push_rels sign env,sign::signs)
else
- (n + p, (List.rev realargs)@l, arsign::signs)
+ (n + p, (List.rev realargs)@l, push_rels sign env,sign::signs)
| c,NotInd _ ->
- (n, l, []::signs) in
- let n, allargs, signs = List.fold_left cook (0, [], []) tomatchs in
- let env = List.fold_right push_rels signs env in
+ (n, l, env, []::signs) in
+ let n, allargs, env, signs = List.fold_left cook (0, [], env, []) tomatchs in
+ let names = List.rev (List.map (List.map pi1) signs) in
let allargs =
- List.map (fun c -> lift n (nf_betadeltaiota env (evars_of isevars) c)) allargs in
+ List.map (fun c -> lift n (nf_betadeltaiota env (Evd.evars_of !isevars) c)) allargs in
let rec build_skeleton env c =
(* Don't put into normal form, it has effects on the synthesis of evars *)
(* let c = whd_betadeltaiota env (evars_of isevars) c in *)
(* We turn all subterms possibly dependent into an evar with maximum ctxt*)
if isEvar c or List.exists (eq_constr c) allargs then
- mkExistential isevars env (loc, CasesType)
+ e_new_evar isevars env ~src:(loc, Evd.CasesType)
+ (Retyping.get_type_of env (Evd.evars_of !isevars) c)
else
- map_constr_with_full_binders push_rel build_skeleton env c in
- List.rev (List.map (List.map pi1) signs), build_skeleton env (lift n c)
+ map_constr_with_full_binders push_rel build_skeleton env c
+ in
+ names, build_skeleton env (lift n c)
(* Here, [pred] is assumed to be in the context built from all *)
(* realargs and terms to match *)
@@ -1629,7 +1589,7 @@ let build_initial_predicate isdep allnames pred =
if dependent (mkRel (nar-n')) pred then pred, 1, 1
else liftn (-1) (nar-n') pred, 0, 1
else pred, 0, 0 in
- let na =
+ let na =
if p=1 then
let na = List.hd names in
if na = Anonymous then
@@ -1637,11 +1597,11 @@ let build_initial_predicate isdep allnames pred =
Name (id_of_string "x") (*Hum*)
else na
else Anonymous in
- PrLetIn ((names',na), buildrec (n'+user_p) pred lnames)
+ PrLetIn ((names',na), buildrec (n'+user_p) pred lnames)
in buildrec 0 pred allnames
let extract_arity_signature env0 tomatchl tmsign =
- let get_one_sign n tm {contents = (na,t)} =
+ let get_one_sign n tm (na,t) =
match tm with
| NotInd (bo,typ) ->
(match t with
@@ -1659,9 +1619,10 @@ let extract_arity_signature env0 tomatchl tmsign =
let nparams = List.length params in
if ind <> ind' then
user_err_loc (loc,"",str "Wrong inductive type");
- if List.length nal <> nparams + nrealargs then
- user_err_loc (loc,"",
- str "Wrong number of arguments for inductive type");
+ let nindargs = nparams + nrealargs in
+ (* Normally done at interning time *)
+ if List.length nal <> nindargs then
+ error_wrong_numarg_inductive_loc loc env0 ind' nindargs;
let parnal,realnal = list_chop nparams nal in
if List.exists ((<>) Anonymous) parnal then
user_err_loc (loc,"",
@@ -1679,101 +1640,74 @@ let extract_arity_signature env0 tomatchl tmsign =
| _ -> assert false
in List.rev (buildrec 0 (tomatchl,tmsign))
+let inh_conv_coerce_to_tycon loc env isevars j tycon =
+ match tycon with
+ | Some p ->
+ let (evd',j) = Coercion.inh_conv_coerce_to loc env !isevars j p in
+ isevars := evd';
+ j
+ | None -> j
+
+
(* Builds the predicate. If the predicate is dependent, its context is
* made of 1+nrealargs assumptions for each matched term in an inductive
* type and 1 assumption for each term not _syntactically_ in an
* inductive type.
- * V7 case: determines whether the multiple case is dependent or not
- * - if its arity is made of nrealargs assumptions for each matched
- * term in an inductive type and nothing for terms not _syntactically_
- * in an inductive type, then it is non dependent
- * - if its arity is made of 1+nrealargs assumptions for each matched
- * term in an inductive type and nothing for terms not _syntactically_
- * in an inductive type, then it is dependent and needs an adjustement
- * to fulfill the criterion above that terms not in an inductive type
- * counts for 1 in the dependent case
-
- * V8 case: each matched terms are independently considered dependent
- * or not
+ * Each matched terms are independently considered dependent or not.
- * A type constraint but no annotation case: it is assumed non dependent
+ * A type constraint but no annotation case: it is assumed non dependent.
*)
let prepare_predicate loc typing_fun isevars env tomatchs sign tycon = function
- (* No type annotation at all *)
- | (None,{contents = None}) ->
+ (* No type annotation *)
+ | None ->
(match tycon with
- | None -> None
- | Some t ->
- let names,pred = prepare_predicate_from_tycon loc false env isevars tomatchs t in
- Some (build_initial_predicate false names pred))
-
- (* v8 style type annotation *)
- | (None,{contents = Some rtntyp}) ->
-
+ | Some (None, t) ->
+ let names,pred =
+ prepare_predicate_from_tycon loc false env isevars tomatchs t
+ in Some (build_initial_predicate false names pred)
+ | _ -> None)
+
+ (* Some type annotation *)
+ | Some rtntyp ->
(* We extract the signature of the arity *)
- let arsigns = extract_arity_signature env tomatchs sign in
- let env = List.fold_right push_rels arsigns env in
- let allnames = List.rev (List.map (List.map pi1) arsigns) in
- let predccl = (typing_fun (mk_tycon (new_Type ())) env rtntyp).uj_val in
- Some (build_initial_predicate true allnames predccl)
-
- (* v7 style type annotation; set the v8 annotation by side effect *)
- | (Some pred,x) ->
- let loc = loc_of_rawconstr pred in
- let dep, n, predj =
- let isevars_copy = evars_of isevars in
- (* We first assume the predicate is non dependent *)
- let ndep_arity = build_expected_arity env isevars false tomatchs in
- try
- false, nb_prod ndep_arity, typing_fun (mk_tycon ndep_arity) env pred
- with PretypeError _ | TypeError _ |
- Stdpp.Exc_located (_,(PretypeError _ | TypeError _)) ->
- evars_reset_evd isevars_copy isevars;
- (* We then assume the predicate is dependent *)
- let dep_arity = build_expected_arity env isevars true tomatchs in
- try
- true, nb_prod dep_arity, typing_fun (mk_tycon dep_arity) env pred
- with PretypeError _ | TypeError _ |
- Stdpp.Exc_located (_,(PretypeError _ | TypeError _)) ->
- evars_reset_evd isevars_copy isevars;
- (* Otherwise we attempt to type it without constraints, possibly *)
- (* failing with an error message; it may also be well-typed *)
- (* but fails to satisfy arity constraints in case_dependent *)
- let predj = typing_fun empty_tycon env pred in
- error_wrong_predicate_arity_loc
- loc env predj.uj_val ndep_arity dep_arity
+ let arsign = extract_arity_signature env tomatchs sign in
+ let env = List.fold_right push_rels arsign env in
+ let allnames = List.rev (List.map (List.map pi1) arsign) in
+ let predcclj = typing_fun (mk_tycon (new_Type ())) env rtntyp in
+ let _ =
+ option_app (fun tycon ->
+ isevars := Coercion.inh_conv_coerces_to loc env !isevars predcclj.uj_val tycon)
+ tycon
in
- let ln,predccl= extract_predicate_conclusion dep tomatchs predj.uj_val in
- set_arity_signature dep n sign tomatchs pred x;
- Some (build_initial_predicate dep ln predccl)
-
+ let predccl = (j_nf_isevar !isevars predcclj).uj_val in
+ Some (build_initial_predicate true allnames predccl)
(**************************************************************************)
(* Main entry of the matching compilation *)
-
-let compile_cases loc (typing_fun,isevars) tycon env (predopt, tomatchl, eqns)=
-
+
+let compile_cases loc (typing_fun, isevars) (tycon : Evarutil.type_constraint) env (predopt, tomatchl, eqns)=
+
(* We build the matrix of patterns and right-hand-side *)
let matx = matx_of_eqns env tomatchl eqns in
-
+
(* We build the vector of terms to match consistently with the *)
(* constructors found in patterns *)
- let rawtms, tmsign = List.split tomatchl in
- let tomatchs = coerce_to_indtype typing_fun isevars env matx rawtms in
-
+ let tomatchs = coerce_to_indtype typing_fun isevars env matx tomatchl in
+
(* We build the elimination predicate if any and check its consistency *)
(* with the type of arguments to match *)
+ let tmsign = List.map snd tomatchl in
let pred = prepare_predicate loc typing_fun isevars env tomatchs tmsign tycon predopt in
-
+
(* We deal with initial aliases *)
let matx = prepare_initial_aliases (known_dependent pred) tomatchs matx in
-
+
(* We push the initial terms to match and push their alias to rhs' envs *)
(* names of aliases will be recovered from patterns (hence Anonymous here) *)
let initial_pushed = List.map (fun tm -> Pushed (tm,[])) tomatchs in
-
+
let pb =
{ env = env;
isevars = isevars;
@@ -1783,12 +1717,10 @@ let compile_cases loc (typing_fun,isevars) tycon env (predopt, tomatchl, eqns)=
mat = matx;
caseloc = loc;
typing_function = typing_fun } in
-
+
let _, j = compile pb in
-
- (* We check for unused patterns *)
- List.iter (check_unused_pattern env) matx;
+ (* We check for unused patterns *)
+ List.iter (check_unused_pattern env) matx;
+ inh_conv_coerce_to_tycon loc env isevars j tycon
+end
- match tycon with
- | Some p -> Coercion.inh_conv_coerce_to loc env isevars j p
- | None -> j
diff --git a/pretyping/cases.mli b/pretyping/cases.mli
index 1d2f9025..5919c42a 100644
--- a/pretyping/cases.mli
+++ b/pretyping/cases.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: cases.mli,v 1.22.2.2 2004/07/16 19:30:43 herbelin Exp $ i*)
+(*i $Id: cases.mli 8654 2006-03-22 15:36:58Z msozeau $ i*)
(*i*)
open Util
@@ -23,6 +23,7 @@ type pattern_matching_error =
| BadPattern of constructor * constr
| BadConstructor of constructor * inductive
| WrongNumargConstructor of constructor * int
+ | WrongNumargInductive of inductive * int
| WrongPredicateArity of constr * constr * constr
| NeedsInversion of constr * constr
| UnusedClause of cases_pattern list
@@ -31,26 +32,23 @@ type pattern_matching_error =
exception PatternMatchingError of env * pattern_matching_error
-(*s Used for old cases in pretyping *)
+val error_wrong_numarg_constructor_loc : loc -> env -> constructor -> int -> 'a
-val branch_scheme :
- env -> evar_defs -> bool -> inductive_family -> constr array
-
-type ml_case_error =
- | MlCaseAbsurd
- | MlCaseDependent
-
-exception NotInferable of ml_case_error
-
-val pred_case_ml : (* raises [NotInferable] if not inferable *)
- env -> evar_map -> bool -> inductive_type -> int * types -> constr
+val error_wrong_numarg_inductive_loc : loc -> env -> inductive -> int -> 'a
(*s Compilation of pattern-matching. *)
-val compile_cases :
- loc -> (type_constraint -> env -> rawconstr -> unsafe_judgment)
- * evar_defs -> type_constraint -> env ->
- (rawconstr option * rawconstr option ref) *
- (rawconstr * (name * (loc * inductive * name list) option) ref) list *
- (loc * identifier list * cases_pattern list * rawconstr) list ->
+module type S = sig
+ val compile_cases :
+ loc ->
+ (type_constraint -> env -> rawconstr -> unsafe_judgment) *
+ evar_defs ref ->
+ type_constraint ->
+ env ->
+ rawconstr option *
+ (rawconstr * (name * (loc * inductive * name list) option)) list *
+ (loc * identifier list * cases_pattern list * rawconstr) list ->
unsafe_judgment
+end
+
+module Cases_F(C : Coercion.S) : S
diff --git a/pretyping/cbv.ml b/pretyping/cbv.ml
index 88f59ded..33166ba8 100644
--- a/pretyping/cbv.ml
+++ b/pretyping/cbv.ml
@@ -6,16 +6,16 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: cbv.ml,v 1.12.2.1 2004/07/16 19:30:44 herbelin Exp $ *)
+(* $Id: cbv.ml 7639 2005-12-02 10:01:15Z gregoire $ *)
open Util
open Pp
open Term
open Names
open Environ
-open Instantiate
open Univ
open Evd
+open Conv_oracle
open Closure
open Esubst
@@ -92,7 +92,7 @@ let contract_cofixp env (i,(_,_,bds as bodies)) =
subst_bodies_from_i 0 env, bds.(i)
let make_constr_ref n = function
- | FarRelKey p -> mkRel (n+p)
+ | RelKey p -> mkRel (n+p)
| VarKey id -> mkVar id
| ConstKey cst -> mkConst cst
@@ -128,7 +128,7 @@ let stack_app appl stack =
open RedFlags
let red_set_ref flags = function
- | FarRelKey _ -> red_set flags fDELTA
+ | RelKey _ -> red_set flags fDELTA
| VarKey id -> red_set flags (fVAR id)
| ConstKey sp -> red_set flags (fCONST sp)
@@ -186,7 +186,7 @@ let rec norm_head info env t stack =
let nargs = Array.map (cbv_stack_term info TOP env) args in
norm_head info env head (stack_app (Array.to_list nargs) stack)
| Case (ci,p,c,v) -> norm_head info env c (CASE(p,v,ci,env,stack))
- | Cast (ct,_) -> norm_head info env ct stack
+ | Cast (ct,_,_) -> norm_head info env ct stack
(* constants, axioms
* the first pattern is CRUCIAL, n=0 happens very often:
@@ -196,7 +196,7 @@ let rec norm_head info env t stack =
| Inl (0,v) -> strip_appl v stack
| Inl (n,v) -> strip_appl (shift_value n v) stack
| Inr (n,None) -> (VAL(0, mkRel n), stack)
- | Inr (n,Some p) -> norm_head_ref (n-p) info env stack (FarRelKey p))
+ | Inr (n,Some p) -> norm_head_ref (n-p) info env stack (RelKey p))
| Var id -> norm_head_ref 0 info env stack (VarKey id)
diff --git a/pretyping/cbv.mli b/pretyping/cbv.mli
index bf8e03b3..dfdf12dd 100644
--- a/pretyping/cbv.mli
+++ b/pretyping/cbv.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: cbv.mli,v 1.6.14.1 2004/07/16 19:30:44 herbelin Exp $ i*)
+(*i $Id: cbv.mli 5920 2004-07-16 20:01:26Z herbelin $ i*)
(*i*)
open Names
diff --git a/pretyping/classops.ml b/pretyping/classops.ml
index 2d8fb951..b6cce031 100755..100644
--- a/pretyping/classops.ml
+++ b/pretyping/classops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: classops.ml,v 1.48.2.1 2004/07/16 19:30:44 herbelin Exp $ *)
+(* $Id: classops.ml 8642 2006-03-17 10:09:02Z notin $ *)
open Util
open Pp
@@ -21,6 +21,7 @@ open Term
open Termops
open Rawterm
open Decl_kinds
+open Mod_subst
(* usage qque peu general: utilise aussi dans record *)
@@ -35,14 +36,14 @@ type cl_typ =
| CL_IND of inductive
type cl_info_typ = {
- cl_strength : strength;
cl_param : int
}
type coe_typ = global_reference
type coe_info_typ = {
- coe_value : unsafe_judgment;
+ coe_value : constr;
+ coe_type : types;
coe_strength : strength;
coe_is_identity : bool;
coe_param : int }
@@ -91,11 +92,7 @@ let unfreeze (fcl,fco,fig) =
(* ajout de nouveaux "objets" *)
let add_new_class cl s =
- try
- let n,s' = Bijint.revmap cl !class_tab in
- if s.cl_strength = Global & s'.cl_strength <> Global then
- class_tab := Bijint.replace n cl s !class_tab
- with Not_found ->
+ if not (Bijint.mem cl !class_tab) then
class_tab := Bijint.add cl s !class_tab
let add_new_coercion coe s =
@@ -106,11 +103,19 @@ let add_new_path x y =
let init () =
class_tab:= Bijint.empty;
- add_new_class CL_FUN { cl_param = 0; cl_strength = Global };
- add_new_class CL_SORT { cl_param = 0; cl_strength = Global };
+ add_new_class CL_FUN { cl_param = 0 };
+ add_new_class CL_SORT { cl_param = 0 };
coercion_tab:= Gmap.empty;
inheritance_graph:= Gmap.empty
+let _ =
+ Summary.declare_summary "inh_graph"
+ { Summary.freeze_function = freeze;
+ Summary.unfreeze_function = unfreeze;
+ Summary.init_function = init;
+ Summary.survive_module = false;
+ Summary.survive_section = false }
+
let _ = init()
(* class_info : cl_typ -> int * cl_info_typ *)
@@ -146,80 +151,44 @@ let lookup_pattern_path_between (s,t) =
(fun coe ->
let c, _ =
Reductionops.whd_betadeltaiota_stack (Global.env()) Evd.empty
- coe.coe_value.uj_val
+ coe.coe_value
in
match kind_of_term c with
| Construct sp -> (sp, coe.coe_param)
| _ -> raise Not_found) l
+(* find_class_type : constr -> cl_typ * int *)
+
+let find_class_type t =
+ let t', args = decompose_app (Reductionops.whd_betaiotazeta t) in
+ match kind_of_term t' with
+ | Var id -> CL_SECVAR id, args
+ | Const sp -> CL_CONST sp, args
+ | Ind ind_sp -> CL_IND ind_sp, args
+ | Prod (_,_,_) -> CL_FUN, []
+ | Sort _ -> CL_SORT, []
+ | _ -> raise Not_found
+
let subst_cl_typ subst ct = match ct with
CL_SORT
| CL_FUN
| CL_SECVAR _ -> ct
| CL_CONST kn ->
- let kn' = subst_kn subst kn in
+ let kn',t = subst_con subst kn in
if kn' == kn then ct else
- CL_CONST kn'
+ fst (find_class_type t)
| CL_IND (kn,i) ->
let kn' = subst_kn subst kn in
if kn' == kn then ct else
CL_IND (kn',i)
-let subst_coe_typ = subst_global
-
-let subst_coe_info subst info =
- let jud = info.coe_value in
- let val' = subst_mps subst (j_val jud) in
- let type' = subst_mps subst (j_type jud) in
- if val' == j_val jud && type' == j_type jud then info else
- {info with coe_value = make_judge val' type'}
-
-(* library, summary *)
-
-(*val inClass : (cl_typ * cl_info_typ) -> Libobject.object = <fun>
- val outClass : Libobject.object -> (cl_typ * cl_info_typ) = <fun> *)
-
-let cache_class (_,(x,y)) = add_new_class x y
-
-let subst_class (_,subst,(ct,ci as obj)) =
- let ct' = subst_cl_typ subst ct in
- if ct' == ct then obj else
- (ct',ci)
-
-let (inClass,outClass) =
- declare_object {(default_object "CLASS") with
- load_function = (fun _ o -> cache_class o);
- cache_function = cache_class;
- subst_function = subst_class;
- classify_function = (fun (_,x) -> Substitute x);
- export_function = (function x -> Some x) }
-
-let declare_class (cl,stre,p) =
- Lib.add_anonymous_leaf (inClass ((cl,{ cl_strength = stre; cl_param = p })))
-
-let _ =
- Summary.declare_summary "inh_graph"
- { Summary.freeze_function = freeze;
- Summary.unfreeze_function = unfreeze;
- Summary.init_function = init;
- Summary.survive_module = false;
- Summary.survive_section = false }
+(*CSC: here we should change the datatype for coercions: it should be possible
+ to declare any term as a coercion *)
+let subst_coe_typ subst t = fst (subst_global subst t)
(* classe d'un terme *)
-(* find_class_type : constr -> cl_typ * int *)
-
-let find_class_type t =
- let t', args = decompose_app (Reductionops.whd_betaiotazeta t) in
- match kind_of_term t' with
- | Var id -> CL_SECVAR id, args
- | Const sp -> CL_CONST sp, args
- | Ind ind_sp -> CL_IND ind_sp, args
- | Prod (_,_,_) -> CL_FUN, []
- | Sort _ -> CL_SORT, []
- | _ -> raise Not_found
-
(* class_of : Term.constr -> int *)
let class_of env sigma t =
@@ -241,8 +210,8 @@ let inductive_class_of ind = fst (class_info (CL_IND ind))
let class_args_of c = snd (decompose_app c)
let string_of_class = function
- | CL_FUN -> if !Options.v7 then "FUNCLASS" else "Funclass"
- | CL_SORT -> if !Options.v7 then "SORTCLASS" else "Sortclass"
+ | CL_FUN -> "Funclass"
+ | CL_SORT -> "Sortclass"
| CL_CONST sp ->
string_of_qualid (shortest_qualid_of_global Idset.empty (ConstRef sp))
| CL_IND sp ->
@@ -254,7 +223,8 @@ let pr_class x = str (string_of_class x)
(* coercion_value : coe_index -> unsafe_judgment * bool *)
-let coercion_value { coe_value = j; coe_is_identity = b } = (j,b)
+let coercion_value { coe_value = c; coe_type = t; coe_is_identity = b } =
+ (make_judge c t, b)
(* pretty-print functions are now in Pretty *)
(* rajouter une coercion dans le graphe *)
@@ -319,49 +289,81 @@ let add_coercion_in_graph (ic,source,target) =
if (!ambig_paths <> []) && is_verbose () then
ppnl (message_ambig !ambig_paths)
-type coercion = coe_typ * coe_info_typ * cl_typ * cl_typ
+type coercion = coe_typ * strength * bool * cl_typ * cl_typ * int
+
+(* Calcul de l'arité d'une classe *)
-let cache_coercion (_,(coe,xf,cls,clt)) =
+let reference_arity_length ref =
+ let t = Global.type_of_global ref in
+ List.length (fst (Reductionops.splay_arity (Global.env()) Evd.empty t))
+
+let class_params = function
+ | CL_FUN | CL_SORT -> 0
+ | CL_CONST sp -> reference_arity_length (ConstRef sp)
+ | CL_SECVAR sp -> reference_arity_length (VarRef sp)
+ | CL_IND sp -> reference_arity_length (IndRef sp)
+
+(* add_class : cl_typ -> strength option -> bool -> unit *)
+
+let add_class cl =
+ add_new_class cl { cl_param = class_params cl }
+
+let load_coercion i (_,(coe,stre,isid,cls,clt,ps)) =
+ add_class cls;
+ add_class clt;
let is,_ = class_info cls in
let it,_ = class_info clt in
+ let xf =
+ { coe_value = constr_of_global coe;
+ coe_type = Global.type_of_global coe;
+ coe_strength = stre;
+ coe_is_identity = isid;
+ coe_param = ps } in
add_new_coercion coe xf;
add_coercion_in_graph (xf,is,it)
-let subst_coercion (_,subst,(coe,xf,cls,clt as obj)) =
+let cache_coercion o =
+ load_coercion 1 o
+
+let subst_coercion (_,subst,(coe,stre,isid,cls,clt,ps as obj)) =
let coe' = subst_coe_typ subst coe in
- let xf' = subst_coe_info subst xf in
let cls' = subst_cl_typ subst cls in
let clt' = subst_cl_typ subst clt in
- if coe' == coe && xf' == xf && cls' == cls & clt' == clt then obj else
- (coe',xf',cls',clt')
-
-
-(* val inCoercion : coercion -> Libobject.object
- val outCoercion : Libobject.object -> coercion *)
+ if coe' == coe && cls' == cls & clt' == clt then obj else
+ (coe',stre,isid,cls',clt',ps)
+
+let discharge_cl = function
+ | CL_CONST kn -> CL_CONST (Lib.discharge_con kn)
+ | CL_IND ind -> CL_IND (Lib.discharge_inductive ind)
+ | cl -> cl
+
+let discharge_coercion (_,(coe,stre,isid,cls,clt,ps)) =
+ if stre = Local then None else
+ let n = try Array.length (Lib.section_instance coe) with Not_found -> 0 in
+ Some (Lib.discharge_global coe,
+ stre,
+ isid,
+ discharge_cl cls,
+ discharge_cl clt,
+ n + ps)
let (inCoercion,outCoercion) =
declare_object {(default_object "COERCION") with
- load_function = (fun _ o -> cache_coercion o);
- cache_function = cache_coercion;
- subst_function = subst_coercion;
- classify_function = (fun (_,x) -> Substitute x);
- export_function = (function x -> Some x) }
-
-let declare_coercion coef v stre ~isid ~src:cls ~target:clt ~params:ps =
- Lib.add_anonymous_leaf
- (inCoercion
- (coef,
- { coe_value = v;
- coe_strength = stre;
- coe_is_identity = isid;
- coe_param = ps },
- cls, clt))
+ load_function = load_coercion;
+ cache_function = cache_coercion;
+ subst_function = subst_coercion;
+ classify_function = (fun (_,x) -> Substitute x);
+ discharge_function = discharge_coercion;
+ export_function = (function x -> Some x) }
+
+let declare_coercion coef stre ~isid ~src:cls ~target:clt ~params:ps =
+ Lib.add_anonymous_leaf (inCoercion (coef,stre,isid,cls,clt,ps))
let coercion_strength v = v.coe_strength
let coercion_identity v = v.coe_is_identity
(* For printing purpose *)
-let get_coercion_value v = v.coe_value.uj_val
+let get_coercion_value v = v.coe_value
let classes () = Bijint.dom !class_tab
let coercions () = Gmap.rng !coercion_tab
diff --git a/pretyping/classops.mli b/pretyping/classops.mli
index f846a9e5..276b14d1 100644
--- a/pretyping/classops.mli
+++ b/pretyping/classops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: classops.mli,v 1.30.2.1 2004/07/16 19:30:44 herbelin Exp $ i*)
+(*i $Id: classops.mli 6748 2005-02-18 22:17:50Z herbelin $ i*)
(*i*)
open Names
@@ -15,6 +15,7 @@ open Term
open Evd
open Environ
open Nametab
+open Mod_subst
(*i*)
(*s This is the type of class kinds *)
@@ -29,7 +30,6 @@ val subst_cl_typ : substitution -> cl_typ -> cl_typ
(* This is the type of infos for declared classes *)
type cl_info_typ = {
- cl_strength : strength;
cl_param : int }
(* This is the type of coercion kinds *)
@@ -47,9 +47,6 @@ type coe_index
(* This is the type of paths from a class to another *)
type inheritance_path = coe_index list
-(*s [declare_class] adds a class to the set of declared classes *)
-val declare_class : cl_typ * strength * int -> unit
-
(*s Access to classes infos *)
val class_info : cl_typ -> (cl_index * cl_info_typ)
val class_exists : cl_typ -> bool
@@ -69,7 +66,7 @@ val class_args_of : constr -> constr list
(*s [declare_coercion] adds a coercion in the graph of coercion paths *)
val declare_coercion :
- coe_typ -> unsafe_judgment -> strength -> isid:bool ->
+ coe_typ -> strength -> isid:bool ->
src:cl_typ -> target:cl_typ -> params:int -> unit
(*s Access to coercions infos *)
@@ -84,19 +81,6 @@ val lookup_path_to_sort_from : cl_index -> inheritance_path
val lookup_pattern_path_between :
cl_index * cl_index -> (constructor * int) list
-(*i Pour le discharge *)
-type coercion = coe_typ * coe_info_typ * cl_typ * cl_typ
-
-open Libobject
-val inClass : (cl_typ * cl_info_typ) -> Libobject.obj
-val outClass : Libobject.obj -> (cl_typ * cl_info_typ)
-val inCoercion : coercion -> Libobject.obj
-val outCoercion : Libobject.obj -> coercion
-val coercion_strength : coe_info_typ -> strength
-val coercion_identity : coe_info_typ -> bool
-val coercion_params : coe_info_typ -> int
-(*i*)
-
(*i Crade *)
open Pp
val install_path_printer :
diff --git a/pretyping/clenv.ml b/pretyping/clenv.ml
new file mode 100644
index 00000000..0b88b14b
--- /dev/null
+++ b/pretyping/clenv.ml
@@ -0,0 +1,435 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id: clenv.ml 8688 2006-04-07 15:08:12Z msozeau $ *)
+
+open Pp
+open Util
+open Names
+open Nameops
+open Term
+open Termops
+open Sign
+open Environ
+open Evd
+open Reduction
+open Reductionops
+open Rawterm
+open Pattern
+open Tacexpr
+open Tacred
+open Pretype_errors
+open Evarutil
+open Unification
+open Mod_subst
+
+(* *)
+let w_coerce env c ctyp target evd =
+ let j = make_judge c ctyp in
+ let (evd',j') = Coercion.Default.inh_conv_coerce_to dummy_loc env evd j (mk_tycon_type target) in
+ (evd',j'.uj_val)
+
+let pf_env gls = Global.env_of_context gls.it.evar_hyps
+let pf_type_of gls c = Typing.type_of (pf_env gls) gls.sigma c
+let pf_hnf_constr gls c = hnf_constr (pf_env gls) gls.sigma c
+let pf_concl gl = gl.it.evar_concl
+
+(******************************************************************)
+(* Clausal environments *)
+
+type clausenv = {
+ templenv : env;
+ env : evar_defs;
+ templval : constr freelisted;
+ templtyp : constr freelisted }
+
+let cl_env ce = ce.templenv
+let cl_sigma ce = evars_of ce.env
+
+let subst_clenv sub clenv =
+ { templval = map_fl (subst_mps sub) clenv.templval;
+ templtyp = map_fl (subst_mps sub) clenv.templtyp;
+ env = subst_evar_defs sub clenv.env;
+ templenv = clenv.templenv }
+
+let clenv_nf_meta clenv c = nf_meta clenv.env c
+let clenv_meta_type clenv mv = Typing.meta_type clenv.env mv
+let clenv_value clenv = meta_instance clenv.env clenv.templval
+let clenv_type clenv = meta_instance clenv.env clenv.templtyp
+
+
+let clenv_hnf_constr ce t = hnf_constr (cl_env ce) (cl_sigma ce) t
+
+let clenv_get_type_of ce c =
+ let metamap =
+ List.map
+ (function
+ | (n,Clval(_,_,typ)) -> (n,typ.rebus)
+ | (n,Cltyp (_,typ)) -> (n,typ.rebus))
+ (meta_list ce.env) in
+ Retyping.get_type_of_with_meta (cl_env ce) (cl_sigma ce) metamap c
+
+let clenv_environments evd bound c =
+ let rec clrec (e,metas) n c =
+ match n, kind_of_term c with
+ | (Some 0, _) -> (e, List.rev metas, c)
+ | (n, Cast (c,_,_)) -> clrec (e,metas) n c
+ | (n, Prod (na,c1,c2)) ->
+ let mv = new_meta () in
+ let dep = dependent (mkRel 1) c2 in
+ let na' = if dep then na else Anonymous in
+ let e' = meta_declare mv c1 ~name:na' e in
+ clrec (e', (mkMeta mv)::metas) (option_app ((+) (-1)) n)
+ (if dep then (subst1 (mkMeta mv) c2) else c2)
+ | (n, LetIn (na,b,_,c)) ->
+ clrec (e,metas) (option_app ((+) (-1)) n) (subst1 b c)
+ | (n, _) -> (e, List.rev metas, c)
+ in
+ clrec (evd,[]) bound c
+
+let clenv_environments_evars env evd bound c =
+ let rec clrec (e,ts) n c =
+ match n, kind_of_term c with
+ | (Some 0, _) -> (e, List.rev ts, c)
+ | (n, Cast (c,_,_)) -> clrec (e,ts) n c
+ | (n, Prod (na,c1,c2)) ->
+ let e',constr = Evarutil.new_evar e env c1 in
+ let dep = dependent (mkRel 1) c2 in
+ clrec (e', constr::ts) (option_app ((+) (-1)) n)
+ (if dep then (subst1 constr c2) else c2)
+ | (n, LetIn (na,b,_,c)) ->
+ clrec (e,ts) (option_app ((+) (-1)) n) (subst1 b c)
+ | (n, _) -> (e, List.rev ts, c)
+ in
+ clrec (evd,[]) bound c
+
+let mk_clenv_from_n gls n (c,cty) =
+ let evd = create_evar_defs gls.sigma in
+ let (env,args,concl) = clenv_environments evd n cty in
+ { templval = mk_freelisted (match args with [] -> c | _ -> applist (c,args));
+ templtyp = mk_freelisted concl;
+ env = env;
+ templenv = Global.env_of_context gls.it.evar_hyps }
+
+let mk_clenv_from gls = mk_clenv_from_n gls None
+
+let mk_clenv_rename_from gls (c,t) =
+ mk_clenv_from gls (c,rename_bound_var (pf_env gls) [] t)
+
+let mk_clenv_rename_from_n gls n (c,t) =
+ mk_clenv_from_n gls n (c,rename_bound_var (pf_env gls) [] t)
+
+let mk_clenv_type_of gls t = mk_clenv_from gls (t,pf_type_of gls t)
+
+(******************************************************************)
+
+(* [mentions clenv mv0 mv1] is true if mv1 is defined and mentions
+ * mv0, or if one of the free vars on mv1's freelist mentions
+ * mv0 *)
+
+let mentions clenv mv0 =
+ let rec menrec mv1 =
+ mv0 = mv1 ||
+ let mlist =
+ try (meta_fvalue clenv.env mv1).freemetas
+ with Anomaly _ | Not_found -> Metaset.empty in
+ meta_exists menrec mlist
+ in menrec
+
+let clenv_defined clenv mv = meta_defined clenv.env mv
+
+let error_incompatible_inst clenv mv =
+ let na = meta_name clenv.env mv in
+ match na with
+ Name id ->
+ errorlabstrm "clenv_assign"
+ (str "An incompatible instantiation has already been found for " ++
+ pr_id id)
+ | _ ->
+ anomaly "clenv_assign: non dependent metavar already assigned"
+
+(* TODO: replace by clenv_unify (mkMeta mv) rhs ? *)
+let clenv_assign mv rhs clenv =
+ let rhs_fls = mk_freelisted rhs in
+ if meta_exists (mentions clenv mv) rhs_fls.freemetas then
+ error "clenv_assign: circularity in unification";
+ try
+ if meta_defined clenv.env mv then
+ if not (eq_constr (meta_fvalue clenv.env mv).rebus rhs) then
+ error_incompatible_inst clenv mv
+ else
+ clenv
+ else {clenv with env = meta_assign mv rhs_fls.rebus clenv.env}
+ with Not_found ->
+ error "clenv_assign: undefined meta"
+
+
+let clenv_wtactic f clenv = {clenv with env = f clenv.env }
+
+
+(* [clenv_dependent hyps_only clenv]
+ * returns a list of the metavars which appear in the template of clenv,
+ * and which are dependent, This is computed by taking the metavars in cval,
+ * in right-to-left order, and collecting the metavars which appear
+ * in their types, and adding in all the metavars appearing in the
+ * type of clenv.
+ * If [hyps_only] then metavariables occurring in the type are _excluded_ *)
+
+(* collects all metavar occurences, in left-to-right order, preserving
+ * repetitions and all. *)
+
+let collect_metas c =
+ let rec collrec acc c =
+ match kind_of_term c with
+ | Meta mv -> mv::acc
+ | _ -> fold_constr collrec acc c
+ in
+ List.rev (collrec [] c)
+
+(* [clenv_metavars clenv mv]
+ * returns a list of the metavars which appear in the type of
+ * the metavar mv. The list is unordered. *)
+
+let clenv_metavars clenv mv = (meta_ftype clenv mv).freemetas
+
+let dependent_metas clenv mvs conclmetas =
+ List.fold_right
+ (fun mv deps ->
+ Metaset.union deps (clenv_metavars clenv.env mv))
+ mvs conclmetas
+
+let clenv_dependent hyps_only clenv =
+ let mvs = collect_metas (clenv_value clenv) in
+ let ctyp_mvs = (mk_freelisted (clenv_type clenv)).freemetas in
+ let deps = dependent_metas clenv mvs ctyp_mvs in
+ List.filter
+ (fun mv -> Metaset.mem mv deps &&
+ not (hyps_only && Metaset.mem mv ctyp_mvs))
+ mvs
+
+let clenv_missing ce = clenv_dependent true ce
+
+(******************************************************************)
+
+let clenv_unify allow_K cv_pb t1 t2 clenv =
+ { clenv with env = w_unify allow_K clenv.templenv cv_pb t1 t2 clenv.env }
+
+let clenv_unique_resolver allow_K clause gl =
+ clenv_unify allow_K CUMUL (clenv_type clause) (pf_concl gl) clause
+
+
+(* [clenv_pose_dependent_evars clenv]
+ * For each dependent evar in the clause-env which does not have a value,
+ * pose a value for it by constructing a fresh evar. We do this in
+ * left-to-right order, so that every evar's type is always closed w.r.t.
+ * metas. *)
+let clenv_pose_dependent_evars clenv =
+ let dep_mvs = clenv_dependent false clenv in
+ List.fold_left
+ (fun clenv mv ->
+ let ty = clenv_meta_type clenv mv in
+ let (evd,evar) = new_evar clenv.env (cl_env clenv) ty in
+ clenv_assign mv evar {clenv with env=evd})
+ clenv
+ dep_mvs
+
+let evar_clenv_unique_resolver clenv gls =
+ clenv_pose_dependent_evars (clenv_unique_resolver false clenv gls)
+
+
+(******************************************************************)
+
+let connect_clenv gls clenv =
+ { clenv with
+ env = evars_reset_evd gls.sigma clenv.env;
+ templenv = Global.env_of_context gls.it.evar_hyps }
+
+(* [clenv_fchain mv clenv clenv']
+ *
+ * Resolves the value of "mv" (which must be undefined) in clenv to be
+ * the template of clenv' be the value "c", applied to "n" fresh
+ * metavars, whose types are chosen by destructing "clf", which should
+ * be a clausale forme generated from the type of "c". The process of
+ * resolution can cause unification of already-existing metavars, and
+ * of the fresh ones which get created. This operation is a composite
+ * of operations which pose new metavars, perform unification on
+ * terms, and make bindings. *)
+let clenv_fchain mv clenv nextclenv =
+ (* Add the metavars of [nextclenv] to [clenv], with their name-environment *)
+ let clenv' =
+ { templval = clenv.templval;
+ templtyp = clenv.templtyp;
+ env = meta_merge clenv.env nextclenv.env;
+ templenv = nextclenv.templenv } in
+ (* unify the type of the template of [nextclenv] with the type of [mv] *)
+ let clenv'' =
+ clenv_unify true CUMUL
+ (clenv_nf_meta clenv' nextclenv.templtyp.rebus)
+ (clenv_meta_type clenv' mv)
+ clenv' in
+ (* assign the metavar *)
+ let clenv''' =
+ clenv_assign mv (clenv_nf_meta clenv' nextclenv.templval.rebus) clenv''
+ in
+ clenv'''
+
+(***************************************************************)
+(* Bindings *)
+
+type arg_bindings = (int * constr) list
+
+(* [clenv_independent clenv]
+ * returns a list of metavariables which appear in the term cval,
+ * and which are not dependent. That is, they do not appear in
+ * the types of other metavars which are in cval, nor in the type
+ * of cval, ctyp. *)
+
+let clenv_independent clenv =
+ let mvs = collect_metas (clenv_value clenv) in
+ let ctyp_mvs = (mk_freelisted (clenv_type clenv)).freemetas in
+ let deps = dependent_metas clenv mvs ctyp_mvs in
+ List.filter (fun mv -> not (Metaset.mem mv deps)) mvs
+
+let meta_of_binder clause loc b t mvs =
+ match b with
+ | NamedHyp s ->
+ if List.exists (fun (_,b',_) -> b=b') t then
+ errorlabstrm "clenv_match_args"
+ (str "The variable " ++ pr_id s ++
+ str " occurs more than once in binding");
+ meta_with_name clause.env s
+ | AnonHyp n ->
+ if List.exists (fun (_,b',_) -> b=b') t then
+ errorlabstrm "clenv_match_args"
+ (str "The position " ++ int n ++
+ str " occurs more than once in binding");
+ try List.nth mvs (n-1)
+ with (Failure _|Invalid_argument _) ->
+ errorlabstrm "clenv_match_args" (str "No such binder")
+
+let error_already_defined b =
+ match b with
+ NamedHyp id ->
+ errorlabstrm "clenv_match_args"
+ (str "Binder name \"" ++ pr_id id ++
+ str"\" already defined with incompatible value")
+ | AnonHyp n ->
+ anomalylabstrm "clenv_match_args"
+ (str "Position " ++ int n ++ str" already defined")
+
+let clenv_match_args s clause =
+ let mvs = clenv_independent clause in
+ let rec matchrec clause = function
+ | [] -> clause
+ | (loc,b,c)::t ->
+ let k = meta_of_binder clause loc b t mvs in
+ if meta_defined clause.env k then
+ if eq_constr (meta_fvalue clause.env k).rebus c then
+ matchrec clause t
+ else error_already_defined b
+ else
+ let k_typ = clenv_hnf_constr clause (clenv_meta_type clause k)
+ (* nf_betaiota was before in type_of - useful to reduce
+ types like (x:A)([x]P u) *)
+ and c_typ =
+ clenv_hnf_constr clause
+ (nf_betaiota (clenv_get_type_of clause c)) in
+ let cl =
+ (* Try to infer some Meta/Evar from the type of [c] *)
+ try clenv_assign k c (clenv_unify true CUMUL c_typ k_typ clause)
+ with e when precatchable_exception e ->
+ (* Try to coerce to the type of [k]; cannot merge with the
+ previous case because Coercion does not handle Meta *)
+ let (_,c') = w_coerce (cl_env clause) c c_typ k_typ clause.env in
+ try clenv_unify true CONV (mkMeta k) c' clause
+ with PretypeError (env,CannotUnify (m,n)) ->
+ Stdpp.raise_with_loc loc
+ (PretypeError (env,CannotUnifyBindingType (m,n)))
+ in matchrec cl t
+ in
+ matchrec clause s
+
+
+let clenv_constrain_with_bindings bl clause =
+ if bl = [] then
+ clause
+ else
+ let all_mvs = collect_metas clause.templval.rebus in
+ let rec matchrec clause = function
+ | [] -> clause
+ | (n,c)::t ->
+ let k =
+ (try
+ if n > 0 then
+ List.nth all_mvs (n-1)
+ else if n < 0 then
+ List.nth (List.rev all_mvs) (-n-1)
+ else error "clenv_constrain_with_bindings"
+ with Failure _ ->
+ errorlabstrm "clenv_constrain_with_bindings"
+ (str"Clause did not have " ++ int n ++ str"-th" ++
+ str" absolute argument")) in
+ let k_typ = nf_betaiota (clenv_meta_type clause k) in
+ let c_typ = nf_betaiota (clenv_get_type_of clause c) in
+ matchrec
+ (clenv_assign k c (clenv_unify true CUMUL c_typ k_typ clause)) t
+ in
+ matchrec clause bl
+
+
+(* not exported: maybe useful ? *)
+let clenv_constrain_dep_args hyps_only clause = function
+ | [] -> clause
+ | mlist ->
+ let occlist = clenv_dependent hyps_only clause in
+ if List.length occlist = List.length mlist then
+ List.fold_left2
+ (fun clenv k c ->
+ try
+ let k_typ =
+ clenv_hnf_constr clause (clenv_meta_type clause k) in
+ let c_typ =
+ clenv_hnf_constr clause (clenv_get_type_of clause c) in
+ (* faire quelque chose avec le sigma retourne ? *)
+ let (_,c') = w_coerce (cl_env clenv) c c_typ k_typ clenv.env in
+ clenv_unify true CONV (mkMeta k) c' clenv
+ with _ ->
+ clenv_unify true CONV (mkMeta k) c clenv)
+ clause occlist mlist
+ else
+ error ("Not the right number of missing arguments (expected "
+ ^(string_of_int (List.length occlist))^")")
+
+let clenv_constrain_missing_args mlist clause =
+ clenv_constrain_dep_args true clause mlist
+
+
+(****************************************************************)
+(* Clausal environment for an application *)
+
+let make_clenv_binding_gen n gls (c,t) = function
+ | ImplicitBindings largs ->
+ let clause = mk_clenv_from_n gls n (c,t) in
+ clenv_constrain_dep_args (n <> None) clause largs
+ | ExplicitBindings lbind ->
+ let clause = mk_clenv_rename_from_n gls n (c,t) in
+ clenv_match_args lbind clause
+ | NoBindings ->
+ mk_clenv_from_n gls n (c,t)
+
+let make_clenv_binding_apply wc n = make_clenv_binding_gen (Some n) wc
+let make_clenv_binding = make_clenv_binding_gen None
+
+(****************************************************************)
+(* Pretty-print *)
+
+let pr_clenv clenv =
+ h 0
+ (str"TEMPL: " ++ print_constr clenv.templval.rebus ++
+ str" : " ++ print_constr clenv.templtyp.rebus ++ fnl () ++
+ pr_evar_defs clenv.env)
diff --git a/pretyping/clenv.mli b/pretyping/clenv.mli
new file mode 100644
index 00000000..f585dfea
--- /dev/null
+++ b/pretyping/clenv.mli
@@ -0,0 +1,117 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: clenv.mli 7659 2005-12-17 21:07:17Z herbelin $ i*)
+
+(*i*)
+open Util
+open Names
+open Term
+open Sign
+open Environ
+open Evd
+open Evarutil
+open Mod_subst
+(*i*)
+
+(***************************************************************)
+(* The Type of Constructions clausale environments. *)
+
+(* [templenv] is the typing context
+ * [env] is the mapping from metavar and evar numbers to their types
+ * and values.
+ * [templval] is the template which we are trying to fill out.
+ * [templtyp] is its type.
+ * [namenv] is a mapping from metavar numbers to names, for
+ * use in instantiating metavars by name.
+ *)
+type clausenv = {
+ templenv : env;
+ env : evar_defs;
+ templval : constr freelisted;
+ templtyp : constr freelisted }
+
+(* Substitution is not applied on templenv (because [subst_clenv] is
+ applied only on hints which typing env is overwritten by the
+ goal env) *)
+val subst_clenv : substitution -> clausenv -> clausenv
+
+(* subject of clenv (instantiated) *)
+val clenv_value : clausenv -> constr
+(* type of clenv (instantiated) *)
+val clenv_type : clausenv -> types
+(* substitute resolved metas *)
+val clenv_nf_meta : clausenv -> constr -> constr
+(* type of a meta in clenv context *)
+val clenv_meta_type : clausenv -> metavariable -> types
+
+val mk_clenv_from : evar_info sigma -> constr * types -> clausenv
+val mk_clenv_from_n :
+ evar_info sigma -> int option -> constr * types -> clausenv
+val mk_clenv_rename_from : evar_info sigma -> constr * types -> clausenv
+val mk_clenv_rename_from_n :
+ evar_info sigma -> int option -> constr * types -> clausenv
+val mk_clenv_type_of : evar_info sigma -> constr -> clausenv
+
+(***************************************************************)
+(* linking of clenvs *)
+
+val connect_clenv : evar_info sigma -> clausenv -> clausenv
+val clenv_fchain : metavariable -> clausenv -> clausenv -> clausenv
+
+(***************************************************************)
+(* Unification with clenvs *)
+
+(* Unifies two terms in a clenv. The boolean is [allow_K] (see [Unification]) *)
+val clenv_unify :
+ bool -> conv_pb -> constr -> constr -> clausenv -> clausenv
+
+(* unifies the concl of the goal with the type of the clenv *)
+val clenv_unique_resolver :
+ bool -> clausenv -> evar_info sigma -> clausenv
+
+(* same as above ([allow_K=false]) but replaces remaining metas
+ with fresh evars *)
+val evar_clenv_unique_resolver :
+ clausenv -> evar_info sigma -> clausenv
+
+(***************************************************************)
+(* Bindings *)
+
+(* bindings where the key is the position in the template of the
+ clenv (dependent or not). Positions can be negative meaning to
+ start from the rightmost argument of the template. *)
+type arg_bindings = (int * constr) list
+
+val clenv_independent : clausenv -> metavariable list
+val clenv_missing : clausenv -> metavariable list
+
+(* defines metas corresponding to the name of the bindings *)
+val clenv_match_args :
+ constr Rawterm.explicit_bindings -> clausenv -> clausenv
+val clenv_constrain_with_bindings : arg_bindings -> clausenv -> clausenv
+
+(* start with a clenv to refine with a given term with bindings *)
+
+(* 1- the arity of the lemma is fixed *)
+val make_clenv_binding_apply :
+ evar_info sigma -> int -> constr * constr -> constr Rawterm.bindings ->
+ clausenv
+val make_clenv_binding :
+ evar_info sigma -> constr * constr -> constr Rawterm.bindings -> clausenv
+
+(* other stuff *)
+val clenv_environments :
+ evar_defs -> int option -> types -> evar_defs * constr list * types
+val clenv_environments_evars :
+ env -> evar_defs -> int option -> types -> evar_defs * constr list * types
+
+(***************************************************************)
+(* Pretty-print *)
+val pr_clenv : clausenv -> Pp.std_ppcmds
+
diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml
index be78eb2c..e01dac47 100644
--- a/pretyping/coercion.ml
+++ b/pretyping/coercion.ml
@@ -5,7 +5,7 @@
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: coercion.ml,v 1.38.6.2 2005/11/29 21:40:52 letouzey Exp $ *)
+(* $Id: coercion.ml 8695 2006-04-10 16:33:52Z msozeau $ *)
open Util
open Names
@@ -19,189 +19,248 @@ open Recordops
open Evarutil
open Evarconv
open Retyping
+open Evd
-(* Typing operations dealing with coercions *)
-
-let class_of1 env sigma t = class_of env sigma (nf_evar sigma t)
-
-(* Here, funj is a coercion therefore already typed in global context *)
-let apply_coercion_args env argl funj =
- let rec apply_rec acc typ = function
- | [] -> { uj_val = applist (j_val funj,argl);
- uj_type = typ }
- | h::restl ->
- (* On devrait pouvoir s'arranger pour qu'on n'ait pas à faire hnf_constr *)
- match kind_of_term (whd_betadeltaiota env Evd.empty typ) with
- | Prod (_,c1,c2) ->
- (* Typage garanti par l'appel à app_coercion*)
- apply_rec (h::acc) (subst1 h c2) restl
- | _ -> anomaly "apply_coercion_args"
- in
- apply_rec [] funj.uj_type argl
-
-exception NoCoercion
-
-(* appliquer le chemin de coercions de patterns p *)
-
-let apply_pattern_coercion loc pat p =
- List.fold_left
- (fun pat (co,n) ->
- let f i = if i<n then Rawterm.PatVar (loc, Anonymous) else pat in
- Rawterm.PatCstr (loc, co, list_tabulate f (n+1), Anonymous))
- pat p
-
-(* raise Not_found if no coercion found *)
-let inh_pattern_coerce_to loc pat ind1 ind2 =
- let i1 = inductive_class_of ind1 in
- let i2 = inductive_class_of ind2 in
- let p = lookup_pattern_path_between (i1,i2) in
- apply_pattern_coercion loc pat p
-
-(* appliquer le chemin de coercions p à hj *)
-
-let apply_coercion env p hj typ_cl =
- try
- fst (List.fold_left
- (fun (ja,typ_cl) i ->
- let fv,isid = coercion_value i in
- let argl = (class_args_of typ_cl)@[ja.uj_val] in
- let jres = apply_coercion_args env argl fv in
- (if isid then
- { uj_val = ja.uj_val; uj_type = jres.uj_type }
- else
- jres),
- jres.uj_type)
- (hj,typ_cl) p)
- with _ -> anomaly "apply_coercion"
-
-let inh_app_fun env isevars j =
- let t = whd_betadeltaiota env (evars_of isevars) j.uj_type in
- match kind_of_term t with
- | Prod (_,_,_) -> j
- | Evar ev when not (is_defined_evar isevars ev) ->
- let t = define_evar_as_arrow isevars ev in
- { uj_val = j.uj_val; uj_type = t }
- | _ ->
- (try
- let t,i1 = class_of1 env (evars_of isevars) j.uj_type in
- let p = lookup_path_to_fun_from i1 in
- apply_coercion env p j t
- with Not_found -> j)
-
-let inh_tosort_force env isevars j =
- try
- let t,i1 = class_of1 env (evars_of isevars) j.uj_type in
- let p = lookup_path_to_sort_from i1 in
- apply_coercion env p j t
- with Not_found ->
- j
-
-let inh_coerce_to_sort env isevars j =
- let typ = whd_betadeltaiota env (evars_of isevars) j.uj_type in
- match kind_of_term typ with
- | Sort s -> { utj_val = j.uj_val; utj_type = s }
- | Evar ev when not (is_defined_evar isevars ev) ->
- let s = define_evar_as_sort isevars ev in
- { utj_val = j.uj_val; utj_type = s }
- | _ ->
- let j1 = inh_tosort_force env isevars j in
- type_judgment env (j_nf_evar (evars_of isevars) j1)
-
-let inh_coerce_to_fail env isevars c1 hj =
- let hj' =
- try
- let t1,i1 = class_of1 env (evars_of isevars) c1 in
- let t2,i2 = class_of1 env (evars_of isevars) hj.uj_type in
- let p = lookup_path_between (i2,i1) in
- apply_coercion env p hj t2
- with Not_found -> raise NoCoercion
- in
- if the_conv_x_leq env isevars hj'.uj_type c1 then
- hj'
- else
- raise NoCoercion
-
-let rec inh_conv_coerce_to_fail env isevars hj c1 =
- let {uj_val = v; uj_type = t} = hj in
- if the_conv_x_leq env isevars t c1 then hj
- else
- try
- inh_coerce_to_fail env isevars c1 hj
- with NoCoercion -> (* try ... with _ -> ... is BAD *)
- (match kind_of_term (whd_betadeltaiota env (evars_of isevars) t),
- kind_of_term (whd_betadeltaiota env (evars_of isevars) c1) with
- | Prod (_,t1,t2), Prod (name,u1,u2) ->
- let v' = whd_betadeltaiota env (evars_of isevars) v in
- if (match kind_of_term v' with
- | Lambda (_,v1,v2) ->
- the_conv_x env isevars v1 u1 (* leq v1 u1? *)
- | _ -> false)
- then
- let (x,v1,v2) = destLambda v' in
- let env1 = push_rel (x,None,v1) env in
- let h2 = inh_conv_coerce_to_fail env1 isevars
- {uj_val = v2; uj_type = t2 } u2 in
- { uj_val = mkLambda (x, v1, h2.uj_val);
- uj_type = mkProd (x, v1, h2.uj_type) }
- else
- (* Mismatch on t1 and u1 or not a lambda: we eta-expand *)
- (* we look for a coercion c:u1->t1 s.t. [name:u1](v' (c x)) *)
- (* has type (name:u1)u2 (with v' recursively obtained) *)
- let name = (match name with
- | Anonymous -> Name (id_of_string "x")
- | _ -> name) in
- let env1 = push_rel (name,None,u1) env in
- let h1 =
- inh_conv_coerce_to_fail env1 isevars
- {uj_val = mkRel 1; uj_type = (lift 1 u1) }
- (lift 1 t1) in
- let h2 = inh_conv_coerce_to_fail env1 isevars
- { uj_val = mkApp (lift 1 v, [|h1.uj_val|]);
- uj_type = subst1 h1.uj_val t2 }
- u2
- in
- { uj_val = mkLambda (name, u1, h2.uj_val);
- uj_type = mkProd (name, u1, h2.uj_type) }
- | _ -> raise NoCoercion)
-
-(* Look for cj' obtained from cj by inserting coercions, s.t. cj'.typ = t *)
-let inh_conv_coerce_to loc env isevars cj t =
- let cj' =
+module type S = sig
+ (*s Coercions. *)
+
+ (* [inh_app_fun env isevars j] coerces [j] to a function; i.e. it
+ inserts a coercion into [j], if needed, in such a way it gets as
+ type a product; it returns [j] if no coercion is applicable *)
+ val inh_app_fun :
+ env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_judgment
+
+ (* [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it
+ inserts a coercion into [j], if needed, in such a way it gets as
+ type a sort; it fails if no coercion is applicable *)
+ val inh_coerce_to_sort : loc ->
+ env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_type_judgment
+
+ (* [inh_conv_coerce_to loc env isevars j t] coerces [j] to an object of type
+ [t]; i.e. it inserts a coercion into [j], if needed, in such a way [t] and
+ [j.uj_type] are convertible; it fails if no coercion is applicable *)
+ val inh_conv_coerce_to : loc ->
+ env -> evar_defs -> unsafe_judgment -> type_constraint_type -> evar_defs * unsafe_judgment
+
+
+ (* [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t]
+ is coercible to an object of type [t'] adding evar constraints if needed;
+ it fails if no coercion exists *)
+ val inh_conv_coerces_to : loc ->
+ env -> evar_defs -> types -> type_constraint_type -> evar_defs
+
+ (* [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases
+ pattern [pat] typed in [ind1] into a pattern typed in [ind2];
+ raises [Not_found] if no coercion found *)
+ val inh_pattern_coerce_to :
+ loc -> Rawterm.cases_pattern -> inductive -> inductive -> Rawterm.cases_pattern
+end
+
+module Default = struct
+ (* Typing operations dealing with coercions *)
+ exception NoCoercion
+
+ let class_of1 env sigma t = class_of env sigma (nf_evar sigma t)
+
+ (* Here, funj is a coercion therefore already typed in global context *)
+ let apply_coercion_args env argl funj =
+ let rec apply_rec acc typ = function
+ | [] -> { uj_val = applist (j_val funj,argl);
+ uj_type = typ }
+ | h::restl ->
+ (* On devrait pouvoir s'arranger pour qu'on n'ait pas à faire hnf_constr *)
+ match kind_of_term (whd_betadeltaiota env Evd.empty typ) with
+ | Prod (_,c1,c2) ->
+ (* Typage garanti par l'appel à app_coercion*)
+ apply_rec (h::acc) (subst1 h c2) restl
+ | _ -> anomaly "apply_coercion_args"
+ in
+ apply_rec [] funj.uj_type argl
+
+ (* appliquer le chemin de coercions de patterns p *)
+
+ let apply_pattern_coercion loc pat p =
+ List.fold_left
+ (fun pat (co,n) ->
+ let f i = if i<n then Rawterm.PatVar (loc, Anonymous) else pat in
+ Rawterm.PatCstr (loc, co, list_tabulate f (n+1), Anonymous))
+ pat p
+
+ (* raise Not_found if no coercion found *)
+ let inh_pattern_coerce_to loc pat ind1 ind2 =
+ let i1 = inductive_class_of ind1 in
+ let i2 = inductive_class_of ind2 in
+ let p = lookup_pattern_path_between (i1,i2) in
+ apply_pattern_coercion loc pat p
+
+ (* appliquer le chemin de coercions p à hj *)
+
+ let apply_coercion env p hj typ_cl =
try
- inh_conv_coerce_to_fail env isevars cj t
- with NoCoercion ->
- let sigma = evars_of isevars in
- error_actual_type_loc loc env sigma cj t
- in
- { uj_val = cj'.uj_val; uj_type = t }
-
-(* [inh_apply_rel_list loc env isevars args f tycon] tries to type [(f
- args)] of type [tycon] (if any) by inserting coercions in front of
- each arg$_i$, if necessary *)
-
-let inh_apply_rel_list apploc env isevars argjl (funloc,funj) tycon =
- let rec apply_rec env n resj = function
- | [] -> resj
- | (loc,hj)::restjl ->
- let sigma = evars_of isevars in
- let resj = inh_app_fun env isevars resj in
- let ntyp = whd_betadeltaiota env sigma resj.uj_type in
- match kind_of_term ntyp with
- | Prod (na,c1,c2) ->
- let hj' =
+ fst (List.fold_left
+ (fun (ja,typ_cl) i ->
+ let fv,isid = coercion_value i in
+ let argl = (class_args_of typ_cl)@[ja.uj_val] in
+ let jres = apply_coercion_args env argl fv in
+ (if isid then
+ { uj_val = ja.uj_val; uj_type = jres.uj_type }
+ else
+ jres),
+ jres.uj_type)
+ (hj,typ_cl) p)
+ with _ -> anomaly "apply_coercion"
+
+ let inh_app_fun env isevars j =
+ let t = whd_betadeltaiota env (evars_of isevars) j.uj_type in
+ match kind_of_term t with
+ | Prod (_,_,_) -> (isevars,j)
+ | Evar ev when not (is_defined_evar isevars ev) ->
+ let (isevars',t) = define_evar_as_arrow isevars ev in
+ (isevars',{ uj_val = j.uj_val; uj_type = t })
+ | _ ->
+ (try
+ let t,i1 = class_of1 env (evars_of isevars) j.uj_type in
+ let p = lookup_path_to_fun_from i1 in
+ (isevars,apply_coercion env p j t)
+ with Not_found -> (isevars,j))
+
+ let inh_tosort_force loc env isevars j =
+ try
+ let t,i1 = class_of1 env (evars_of isevars) j.uj_type in
+ let p = lookup_path_to_sort_from i1 in
+ let j1 = apply_coercion env p j t in
+ (isevars,type_judgment env (j_nf_evar (evars_of isevars) j1))
+ with Not_found ->
+ error_not_a_type_loc loc env (evars_of isevars) j
+
+ let inh_coerce_to_sort loc env isevars j =
+ let typ = whd_betadeltaiota env (evars_of isevars) j.uj_type in
+ match kind_of_term typ with
+ | Sort s -> (isevars,{ utj_val = j.uj_val; utj_type = s })
+ | Evar ev when not (is_defined_evar isevars ev) ->
+ let (isevars',s) = define_evar_as_sort isevars ev in
+ (isevars',{ utj_val = j.uj_val; utj_type = s })
+ | _ ->
+ inh_tosort_force loc env isevars j
+
+ let inh_coerce_to_fail env isevars c1 v t =
+ let v', t' =
+ try
+ let t1,i1 = class_of1 env (evars_of isevars) c1 in
+ let t2,i2 = class_of1 env (evars_of isevars) t in
+ let p = lookup_path_between (i2,i1) in
+ match v with
+ Some v ->
+ let j = apply_coercion env p {uj_val = v; uj_type = t} t2 in
+ Some j.uj_val, j.uj_type
+ | None -> None, t
+ with Not_found -> raise NoCoercion
+ in
+ try (the_conv_x_leq env t' c1 isevars, v', t')
+ with Reduction.NotConvertible -> raise NoCoercion
+ open Pp
+ let rec inh_conv_coerce_to_fail loc env isevars v t c1 =
+ try (the_conv_x_leq env t c1 isevars, v, t)
+ with Reduction.NotConvertible ->
+ (try
+ inh_coerce_to_fail env isevars c1 v t
+ with NoCoercion ->
+ (match kind_of_term (whd_betadeltaiota env (evars_of isevars) t),
+ kind_of_term (whd_betadeltaiota env (evars_of isevars) c1) with
+ | Prod (_,t1,t2), Prod (name,u1,u2) ->
+ let v' = option_app (whd_betadeltaiota env (evars_of isevars)) v in
+ let (evd',b) =
+ match v' with
+ Some v' ->
+ (match kind_of_term v' with
+ | Lambda (x,v1,v2) ->
+ (try the_conv_x env v1 u1 isevars, Some (x, v1, v2) (* leq v1 u1? *)
+ with Reduction.NotConvertible -> (isevars, None))
+ | _ -> (isevars, None))
+ | None -> (isevars, None)
+ in
+ (match b with
+ Some (x, v1, v2) ->
+ let env1 = push_rel (x,None,v1) env in
+ let (evd'', v2', t2') = inh_conv_coerce_to_fail loc env1 evd'
+ (Some v2) t2 u2 in
+ (evd'', option_app (fun v2' -> mkLambda (x, v1, v2')) v2',
+ mkProd (x, v1, t2'))
+ | None ->
+ (* Mismatch on t1 and u1 or not a lambda: we eta-expand *)
+ (* we look for a coercion c:u1->t1 s.t. [name:u1](v' (c x)) *)
+ (* has type (name:u1)u2 (with v' recursively obtained) *)
+ let name = (match name with
+ | Anonymous -> Name (id_of_string "x")
+ | _ -> name) in
+ let env1 = push_rel (name,None,u1) env in
+ let (evd', v1', t1') =
+ inh_conv_coerce_to_fail loc env1 isevars
+ (Some (mkRel 1)) (lift 1 u1) (lift 1 t1)
+ in
+ let (evd'', v2', t2') =
+ let v2 =
+ match v with
+ Some v -> option_app (fun v1' -> mkApp (lift 1 v, [|v1'|])) v1'
+ | None -> None
+ and evd', t2 =
+ match v1' with
+ Some v1' -> evd', subst1 v1' t2
+ | None ->
+ let evd', ev = new_evar evd' env ~src:(loc, InternalHole) t1' in
+ evd', subst1 ev t2
+ in
+ inh_conv_coerce_to_fail loc env1 evd' v2 t2 u2
+ in
+ (evd'', option_app (fun v2' -> mkLambda (name, u1, v2')) v2',
+ mkProd (name, u1, t2')))
+ | _ -> raise NoCoercion))
+
+ (* Look for cj' obtained from cj by inserting coercions, s.t. cj'.typ = t *)
+ let inh_conv_coerce_to loc env isevars cj (n, t) =
+ match n with
+ None ->
+ let (evd', val', type') =
+ try
+ inh_conv_coerce_to_fail loc env isevars (Some cj.uj_val) cj.uj_type t
+ with NoCoercion ->
+ let sigma = evars_of isevars in
+ error_actual_type_loc loc env sigma cj t
+ in
+ let val' = match val' with Some v -> v | None -> assert(false) in
+ let nf = nf_isevar evd' in
+ (evd',{ uj_val = nf val'; uj_type = nf t })
+ | Some (init, cur) -> (isevars, cj)
+
+ let inh_conv_coerces_to loc env (isevars : evar_defs) t (abs, t') = isevars
+ (* Still problematic, as it changes unification
+ let nabsinit, nabs =
+ match abs with
+ None -> 0, 0
+ | Some (init, cur) -> init, cur
+ in
+ try
+ let (rels, rng) =
+ (* a little more effort to get products is needed *)
+ try decompose_prod_n nabs t
+ with _ ->
+ if !Options.debug then
+ msg_warning (str "decompose_prod_n failed");
+ raise (Invalid_argument "Coercion.inh_conv_coerces_to")
+ in
+ (* The final range free variables must have been replaced by evars, we accept only that evars
+ in rng are applied to free vars. *)
+ if noccur_with_meta 0 (succ nabsinit) rng then (
+ let env', t, t' =
+ let env' = List.fold_right (fun (n, t) env -> push_rel (n, None, t) env) rels env in
+ env', rng, lift nabs t'
+ in
try
- inh_conv_coerce_to_fail env isevars hj c1
+ pi1 (inh_conv_coerce_to_fail loc env' isevars None t t')
with NoCoercion ->
- error_cant_apply_bad_type_loc apploc env sigma
- (1,c1,hj.uj_type) resj (List.map snd restjl) in
- let newresj =
- { uj_val = applist (j_val resj, [j_val hj']);
- uj_type = subst1 hj'.uj_val c2 } in
- apply_rec (push_rel (na,None,c1) env) (n+1) newresj restjl
- | _ ->
- error_cant_apply_not_functional_loc
- (join_loc funloc loc) env sigma resj
- (List.map snd restjl)
- in
- apply_rec env 1 funj argjl
-
+ isevars) (* Maybe not enough information to unify *)
+ (*let sigma = evars_of isevars in
+ error_cannot_coerce env' sigma (t, t'))*)
+ else isevars
+ with Invalid_argument _ -> isevars *)
+end
diff --git a/pretyping/coercion.mli b/pretyping/coercion.mli
index 658844eb..f675beff 100644
--- a/pretyping/coercion.mli
+++ b/pretyping/coercion.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: coercion.mli,v 1.20.14.2 2004/07/16 19:30:44 herbelin Exp $ i*)
+(*i $Id: coercion.mli 8688 2006-04-07 15:08:12Z msozeau $ i*)
(*i*)
open Util
@@ -19,28 +19,38 @@ open Evarutil
open Rawterm
(*i*)
-(*s Coercions. *)
+module type S = sig
+ (*s Coercions. *)
+
+ (* [inh_app_fun env isevars j] coerces [j] to a function; i.e. it
+ inserts a coercion into [j], if needed, in such a way it gets as
+ type a product; it returns [j] if no coercion is applicable *)
+ val inh_app_fun :
+ env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_judgment
+
+ (* [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it
+ inserts a coercion into [j], if needed, in such a way it gets as
+ type a sort; it fails if no coercion is applicable *)
+ val inh_coerce_to_sort : loc ->
+ env -> evar_defs -> unsafe_judgment -> evar_defs * unsafe_type_judgment
+
+ (* [inh_conv_coerce_to loc env isevars j t] coerces [j] to an object of type
+ [t]; i.e. it inserts a coercion into [j], if needed, in such a way [t] and
+ [j.uj_type] are convertible; it fails if no coercion is applicable *)
+ val inh_conv_coerce_to : loc ->
+ env -> evar_defs -> unsafe_judgment -> type_constraint_type -> evar_defs * unsafe_judgment
-(* [inh_app_fun env isevars j] coerces [j] to a function; i.e. it
- inserts a coercion into [j], if needed, in such a way it gets as
- type a product; it returns [j] if no coercion is applicable *)
-val inh_app_fun :
- env -> evar_defs -> unsafe_judgment -> unsafe_judgment
+ (* [inh_conv_coerces_to loc env isevars t t'] checks if an object of type [t]
+ is coercible to an object of type [t'] adding evar constraints if needed;
+ it fails if no coercion exists *)
+ val inh_conv_coerces_to : loc ->
+ env -> evar_defs -> types -> type_constraint_type -> evar_defs
+
+ (* [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases
+ pattern [pat] typed in [ind1] into a pattern typed in [ind2];
+ raises [Not_found] if no coercion found *)
+ val inh_pattern_coerce_to :
+ loc -> cases_pattern -> inductive -> inductive -> cases_pattern
+end
-(* [inh_coerce_to_sort env isevars j] coerces [j] to a type; i.e. it
- inserts a coercion into [j], if needed, in such a way it gets as
- type a sort; it fails if no coercion is applicable *)
-val inh_coerce_to_sort :
- env -> evar_defs -> unsafe_judgment -> unsafe_type_judgment
-
-(* [inh_conv_coerce_to loc env isevars j t] coerces [j] to an object of type
- [t]; i.e. it inserts a coercion into [j], if needed, in such a way [t] and
- [j.uj_type] are convertible; it fails if no coercion is applicable *)
-val inh_conv_coerce_to : loc ->
- env -> evar_defs -> unsafe_judgment -> constr -> unsafe_judgment
-
-(* [inh_pattern_coerce_to loc env isevars pat ind1 ind2] coerces the Cases
- pattern [pat] typed in [ind1] into a pattern typed in [ind2];
- raises [Not_found] if no coercion found *)
-val inh_pattern_coerce_to :
- loc -> cases_pattern -> inductive -> inductive -> cases_pattern
+module Default : S
diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml
index 040a185e..3f2aed34 100644
--- a/pretyping/detyping.ml
+++ b/pretyping/detyping.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: detyping.ml,v 1.75.2.5 2005/09/06 14:30:41 herbelin Exp $ *)
+(* $Id: detyping.ml 8624 2006-03-13 17:38:17Z msozeau $ *)
open Pp
open Util
@@ -23,6 +23,10 @@ open Nameops
open Termops
open Libnames
open Nametab
+open Evd
+open Mod_subst
+
+let dl = dummy_loc
(****************************************************************************)
(* Tools for printing of Cases *)
@@ -130,11 +134,21 @@ let synthetize_type () = !synth_type_value
let _ = declare_bool_option
{ optsync = true;
- optname = "synthesizability";
+ optname = "pattern matching return type synthesizability";
optkey = SecondaryTable ("Printing","Synth");
optread = synthetize_type;
optwrite = (:=) synth_type_value }
+let reverse_matching_value = ref true
+let reverse_matching () = !reverse_matching_value
+
+let _ = declare_bool_option
+ { optsync = true;
+ optname = "pattern-matching reversibility";
+ optkey = SecondaryTable ("Printing","Matching");
+ optread = reverse_matching;
+ optwrite = (:=) reverse_matching_value }
+
(* Auxiliary function for MutCase printing *)
(* [computable] tries to tell if the predicate typing the result is inferable*)
@@ -175,7 +189,7 @@ let lookup_name_as_renamed env t s =
if id=s then (Some n)
else lookup avoid' (add_name (Name id) env_names) (n+1) c'
| (Anonymous,avoid') -> lookup avoid' env_names (n+1) (pop c'))
- | Cast (c,_) -> lookup avoid env_names n c
+ | Cast (c,_,_) -> lookup avoid env_names n c
| _ -> None
in lookup (ids_of_named_context (named_context env)) empty_names_context 1 t
@@ -189,10 +203,65 @@ let lookup_index_as_renamed env t n =
(match concrete_name true [] empty_names_context name c' with
| (Name _,_) -> lookup n (d+1) c'
| (Anonymous,_) -> if n=1 then Some d else lookup (n-1) (d+1) c')
- | Cast (c,_) -> lookup n d c
+ | Cast (c,_,_) -> lookup n d c
| _ -> None
in lookup n 1 t
+(**********************************************************************)
+(* Fragile algorithm to reverse pattern-matching compilation *)
+
+let update_name na ((_,e),c) =
+ match na with
+ | Name _ when force_wildcard () & noccurn (list_index na e) c ->
+ Anonymous
+ | _ ->
+ na
+
+let rec decomp_branch n nal b (avoid,env as e) c =
+ if n=0 then (List.rev nal,(e,c))
+ else
+ let na,c,f =
+ match kind_of_term (strip_outer_cast c) with
+ | Lambda (na,_,c) -> na,c,concrete_let_name
+ | LetIn (na,_,_,c) -> na,c,concrete_name
+ | _ ->
+ Name (id_of_string "x"),(applist (lift 1 c, [mkRel 1])),
+ concrete_name in
+ let na',avoid' = f b avoid env na c in
+ decomp_branch (n-1) (na'::nal) b (avoid',add_name na' env) c
+
+let rec build_tree na isgoal e ci cl =
+ let mkpat n rhs pl = PatCstr(dl,(ci.ci_ind,n+1),pl,update_name na rhs) in
+ let cnl = ci.ci_cstr_nargs in
+ List.flatten
+ (list_tabulate (fun i -> contract_branch isgoal e (cnl.(i),mkpat i,cl.(i)))
+ (Array.length cl))
+
+and align_tree nal isgoal (e,c as rhs) = match nal with
+ | [] -> [[],rhs]
+ | na::nal ->
+ match kind_of_term c with
+ | Case (ci,_,c,cl) when c = mkRel (list_index na (snd e)) ->
+ let clauses = build_tree na isgoal e ci cl in
+ List.flatten
+ (List.map (fun (pat,rhs) ->
+ let lines = align_tree nal isgoal rhs in
+ List.map (fun (hd,rest) -> pat::hd,rest) lines)
+ clauses)
+ | _ ->
+ let pat = PatVar(dl,update_name na rhs) in
+ let mat = align_tree nal isgoal rhs in
+ List.map (fun (hd,rest) -> pat::hd,rest) mat
+
+and contract_branch isgoal e (cn,mkpat,b) =
+ let nal,rhs = decomp_branch cn [] isgoal e b in
+ let mat = align_tree nal isgoal rhs in
+ List.map (fun (hd,rhs) -> (mkpat rhs hd,rhs)) mat
+
+(**********************************************************************)
+(* Transform internal representation of pattern-matching into list of *)
+(* clauses *)
+
let is_nondep_branch c n =
try
let _,ccl = decompose_lam_n_assum n c in
@@ -208,26 +277,17 @@ let extract_nondep_branches test c b n =
| _ -> assert false in
if test c n then Some (strip n b) else None
-let detype_case computable detype detype_eqn testdep
- tenv avoid indsp st p k c bl =
+let detype_case computable detype detype_eqns testdep avoid data p c bl =
+ let (indsp,st,nparams,consnargsl,k) = data in
let synth_type = synthetize_type () in
let tomatch = detype c in
-
- (* Find constructors arity *)
- let (mib,mip) = Inductive.lookup_mind_specif tenv indsp in
- let get_consnarg j =
- let typi = mis_nf_constructor_type (indsp,mib,mip) (j+1) in
- let _,t = decompose_prod_n_assum (List.length mip.mind_params_ctxt) typi in
- List.rev (fst (decompose_prod_assum t)) in
- let consnargs = Array.init (Array.length mip.mind_consnames) get_consnarg in
- let consnargsl = Array.map List.length consnargs in
- let alias, aliastyp, newpred, pred =
- if (not !Options.raw_print) & synth_type & computable & bl <> [||] then
- Anonymous, None, None, None
+ let alias, aliastyp, pred=
+ if (not !Options.raw_print) & synth_type & computable & Array.length bl<>0
+ then
+ Anonymous, None, None
else
- let p = option_app detype p in
- match p with
- | None -> Anonymous, None, None, None
+ match option_app detype p with
+ | None -> Anonymous, None, None
| Some p ->
let decompose_lam k c =
let rec lamdec_rec l avoid k c =
@@ -237,10 +297,10 @@ let detype_case computable detype detype_eqn testdep
| c ->
let x = next_ident_away (id_of_string "x") avoid in
lamdec_rec ((Name x)::l) (x::avoid) (k-1)
- (let a = RVar (dummy_loc,x) in
+ (let a = RVar (dl,x) in
match c with
| RApp (loc,p,l) -> RApp (loc,p,l@[a])
- | _ -> (RApp (dummy_loc,c,[a])))
+ | _ -> (RApp (dl,c,[a])))
in
lamdec_rec [] [] k c in
let nl,typ = decompose_lam k p in
@@ -250,13 +310,12 @@ let detype_case computable detype detype_eqn testdep
let aliastyp =
if List.for_all ((=) Anonymous) nl then None
else
- let pars = list_tabulate (fun _ -> Anonymous) mip.mind_nparams
- in Some (dummy_loc,indsp,pars@nl) in
- n, aliastyp, Some typ, Some p
+ let pars = list_tabulate (fun _ -> Anonymous) nparams in
+ Some (dl,indsp,pars@nl) in
+ n, aliastyp, Some typ
in
let constructs = Array.init (Array.length bl) (fun i -> (indsp,i+1)) in
- let eqnv = array_map3 detype_eqn constructs consnargsl bl in
- let eqnl = Array.to_list eqnv in
+ let eqnl = detype_eqns constructs consnargsl bl in
let tag =
try
if !Options.raw_print then
@@ -268,12 +327,10 @@ let detype_case computable detype detype_eqn testdep
else
st
with Not_found -> st
- in
- if tag = RegularStyle then
- RCases (dummy_loc,(pred,ref newpred),[tomatch,ref (alias,aliastyp)],eqnl)
- else
- let bl' = Array.map detype bl in
- if not !Options.v7 && tag = LetStyle && aliastyp = None then
+ in
+ match tag with
+ | LetStyle when aliastyp = None ->
+ let bl' = Array.map detype bl in
let rec decomp_lam_force n avoid l p =
if n = 0 then (List.rev l,p) else
match p with
@@ -285,91 +342,75 @@ let detype_case computable detype detype_eqn testdep
let x = Nameops.next_ident_away (id_of_string "x") avoid in
decomp_lam_force (n-1) (x::avoid) (Name x :: l)
(* eta-expansion *)
- (let a = RVar (dummy_loc,x) in
+ (let a = RVar (dl,x) in
match p with
| RApp (loc,p,l) -> RApp (loc,p,l@[a])
- | _ -> (RApp (dummy_loc,p,[a]))) in
+ | _ -> (RApp (dl,p,[a]))) in
let (nal,d) = decomp_lam_force consnargsl.(0) avoid [] bl'.(0) in
- RLetTuple (dummy_loc,nal,(alias,newpred),tomatch,d)
- else
- let nondepbrs =
- array_map3 (extract_nondep_branches testdep) bl bl' consnargsl in
- if not !Options.v7 && tag = IfStyle && aliastyp = None
- && array_for_all ((<>) None) nondepbrs then
- RIf (dummy_loc,tomatch,(alias,newpred),
- out_some nondepbrs.(0),out_some nondepbrs.(1))
- else if !Options.v7 then
- let rec remove_type avoid args c =
- match c,args with
- | RLambda (loc,na,t,c), _::args ->
- let h = RHole (dummy_loc,BinderType na) in
- RLambda (loc,na,h,remove_type avoid args c)
- | RLetIn (loc,na,b,c), _::args ->
- RLetIn (loc,na,b,remove_type avoid args c)
- | c, (na,None,t)::args ->
- let id = next_name_away_with_default "x" na avoid in
- let h = RHole (dummy_loc,BinderType na) in
- let c = remove_type (id::avoid) args
- (RApp (dummy_loc,c,[RVar (dummy_loc,id)])) in
- RLambda (dummy_loc,Name id,h,c)
- | c, (na,Some b,t)::args ->
- let h = RHole (dummy_loc,BinderType na) in
- let avoid = name_fold (fun x l -> x::l) na avoid in
- RLetIn (dummy_loc,na,h,remove_type avoid args c)
- | c, [] -> c in
- let bl' = array_map2 (remove_type avoid) consnargs bl' in
- ROrderedCase (dummy_loc,tag,pred,tomatch,bl',ref None)
- else
- RCases(dummy_loc,(pred,ref newpred),[tomatch,ref (alias,aliastyp)],eqnl)
-
-
-let rec detype tenv avoid env t =
+ RLetTuple (dl,nal,(alias,pred),tomatch,d)
+ | IfStyle when aliastyp = None ->
+ let bl' = Array.map detype bl in
+ let nondepbrs =
+ array_map3 (extract_nondep_branches testdep) bl bl' consnargsl in
+ if array_for_all ((<>) None) nondepbrs then
+ RIf (dl,tomatch,(alias,pred),
+ out_some nondepbrs.(0),out_some nondepbrs.(1))
+ else
+ RCases (dl,pred,[tomatch,(alias,aliastyp)],eqnl)
+ | _ ->
+ RCases (dl,pred,[tomatch,(alias,aliastyp)],eqnl)
+
+(**********************************************************************)
+(* Main detyping function *)
+
+let rec detype isgoal avoid env t =
match kind_of_term (collapse_appl t) with
| Rel n ->
(try match lookup_name_of_rel n env with
- | Name id -> RVar (dummy_loc, id)
+ | Name id -> RVar (dl, id)
| Anonymous -> anomaly "detype: index to an anonymous variable"
with Not_found ->
let s = "_UNBOUND_REL_"^(string_of_int n)
- in RVar (dummy_loc, id_of_string s))
+ in RVar (dl, id_of_string s))
| Meta n ->
(* Meta in constr are not user-parsable and are mapped to Evar *)
- REvar (dummy_loc, n, None)
+ REvar (dl, n, None)
| Var id ->
(try
- let _ = Global.lookup_named id in RRef (dummy_loc, VarRef id)
+ let _ = Global.lookup_named id in RRef (dl, VarRef id)
with _ ->
- RVar (dummy_loc, id))
- | Sort (Prop c) -> RSort (dummy_loc,RProp c)
- | Sort (Type u) -> RSort (dummy_loc,RType (Some u))
- | Cast (c1,c2) ->
- RCast(dummy_loc,detype tenv avoid env c1,
- detype tenv avoid env c2)
- | Prod (na,ty,c) -> detype_binder tenv BProd avoid env na ty c
- | Lambda (na,ty,c) -> detype_binder tenv BLambda avoid env na ty c
- | LetIn (na,b,_,c) -> detype_binder tenv BLetIn avoid env na b c
+ RVar (dl, id))
+ | Sort (Prop c) -> RSort (dl,RProp c)
+ | Sort (Type u) -> RSort (dl,RType (Some u))
+ | Cast (c1,k,c2) ->
+ RCast(dl,detype isgoal avoid env c1, k,
+ detype isgoal avoid env c2)
+ | Prod (na,ty,c) -> detype_binder isgoal BProd avoid env na ty c
+ | Lambda (na,ty,c) -> detype_binder isgoal BLambda avoid env na ty c
+ | LetIn (na,b,_,c) -> detype_binder isgoal BLetIn avoid env na b c
| App (f,args) ->
- RApp (dummy_loc,detype tenv avoid env f,
- array_map_to_list (detype tenv avoid env) args)
- | Const sp -> RRef (dummy_loc, ConstRef sp)
+ RApp (dl,detype isgoal avoid env f,
+ array_map_to_list (detype isgoal avoid env) args)
+ | Const sp -> RRef (dl, ConstRef sp)
| Evar (ev,cl) ->
- REvar (dummy_loc, ev,
- Some (List.map (detype tenv avoid env) (Array.to_list cl)))
+ REvar (dl, ev,
+ Some (List.map (detype isgoal avoid env) (Array.to_list cl)))
| Ind ind_sp ->
- RRef (dummy_loc, IndRef ind_sp)
+ RRef (dl, IndRef ind_sp)
| Construct cstr_sp ->
- RRef (dummy_loc, ConstructRef cstr_sp)
- | Case (annot,p,c,bl) ->
- let comp = computable p (annot.ci_pp_info.ind_nargs) in
- let ind = annot.ci_ind in
- let st = annot.ci_pp_info.style in
- detype_case comp (detype tenv avoid env) (detype_eqn tenv avoid env)
- is_nondep_branch
- (snd tenv) avoid ind st (Some p) annot.ci_pp_info.ind_nargs c bl
- | Fix (nvn,recdef) -> detype_fix tenv avoid env nvn recdef
- | CoFix (n,recdef) -> detype_cofix tenv avoid env n recdef
-
-and detype_fix tenv avoid env (vn,_ as nvn) (names,tys,bodies) =
+ RRef (dl, ConstructRef cstr_sp)
+ | Case (ci,p,c,bl) ->
+ let comp = computable p (ci.ci_pp_info.ind_nargs) in
+ detype_case comp (detype isgoal avoid env)
+ (detype_eqns isgoal avoid env ci comp)
+ is_nondep_branch avoid
+ (ci.ci_ind,ci.ci_pp_info.style,ci.ci_npar,
+ ci.ci_cstr_nargs,ci.ci_pp_info.ind_nargs)
+ (Some p) c bl
+ | Fix (nvn,recdef) -> detype_fix isgoal avoid env nvn recdef
+ | CoFix (n,recdef) -> detype_cofix isgoal avoid env n recdef
+
+and detype_fix isgoal avoid env (vn,_ as nvn) (names,tys,bodies) =
let def_avoid, def_env, lfi =
Array.fold_left
(fun (avoid, env, l) na ->
@@ -378,14 +419,14 @@ and detype_fix tenv avoid env (vn,_ as nvn) (names,tys,bodies) =
(avoid, env, []) names in
let n = Array.length tys in
let v = array_map3
- (fun c t i -> share_names tenv (i+1) [] def_avoid def_env c (lift n t))
+ (fun c t i -> share_names isgoal (i+1) [] def_avoid def_env c (lift n t))
bodies tys vn in
- RRec(dummy_loc,RFix nvn,Array.of_list (List.rev lfi),
+ RRec(dl,RFix (Array.map (fun i -> i, RStructRec) (fst nvn), snd nvn),Array.of_list (List.rev lfi),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
-and detype_cofix tenv avoid env n (names,tys,bodies) =
+and detype_cofix isgoal avoid env n (names,tys,bodies) =
let def_avoid, def_env, lfi =
Array.fold_left
(fun (avoid, env, l) na ->
@@ -394,19 +435,14 @@ and detype_cofix tenv avoid env n (names,tys,bodies) =
(avoid, env, []) names in
let ntys = Array.length tys in
let v = array_map2
- (fun c t -> share_names tenv 0 [] def_avoid def_env c (lift ntys t))
+ (fun c t -> share_names isgoal 0 [] def_avoid def_env c (lift ntys t))
bodies tys in
- RRec(dummy_loc,RCoFix n,Array.of_list (List.rev lfi),
+ RRec(dl,RCoFix n,Array.of_list (List.rev lfi),
Array.map (fun (bl,_,_) -> bl) v,
Array.map (fun (_,_,ty) -> ty) v,
Array.map (fun (_,bd,_) -> bd) v)
-and share_names tenv n l avoid env c t =
- if !Options.v7 && n=0 then
- let c = detype tenv avoid env c in
- let t = detype tenv avoid env t in
- (List.rev l,c,t)
- else
+and share_names isgoal n l avoid env c t =
match kind_of_term c, kind_of_term t with
(* factorize even when not necessary to have better presentation *)
| Lambda (na,t,c), Prod (na',t',c') ->
@@ -414,47 +450,57 @@ and share_names tenv n l avoid env c t =
Name _, _ -> na
| _, Name _ -> na'
| _ -> na in
- let t = detype tenv avoid env t in
+ let t = detype isgoal avoid env t in
let id = next_name_away na avoid in
let avoid = id::avoid and env = add_name (Name id) env in
- share_names tenv (n-1) ((Name id,None,t)::l) avoid env c c'
+ share_names isgoal (n-1) ((Name id,None,t)::l) avoid env c c'
(* May occur for fix built interactively *)
| LetIn (na,b,t',c), _ when n > 0 ->
- let t' = detype tenv avoid env t' in
- let b = detype tenv avoid env b in
+ let t' = detype isgoal avoid env t' in
+ let b = detype isgoal avoid env b in
let id = next_name_away na avoid in
let avoid = id::avoid and env = add_name (Name id) env in
- share_names tenv n ((Name id,Some b,t')::l) avoid env c t
+ share_names isgoal n ((Name id,Some b,t')::l) avoid env c t
(* Only if built with the f/n notation or w/o let-expansion in types *)
| _, LetIn (_,b,_,t) when n > 0 ->
- share_names tenv n l avoid env c (subst1 b t)
+ share_names isgoal n l avoid env c (subst1 b t)
(* If it is an open proof: we cheat and eta-expand *)
| _, Prod (na',t',c') when n > 0 ->
- let t' = detype tenv avoid env t' in
+ let t' = detype isgoal avoid env t' in
let id = next_name_away na' avoid in
let avoid = id::avoid and env = add_name (Name id) env in
let appc = mkApp (lift 1 c,[|mkRel 1|]) in
- share_names tenv (n-1) ((Name id,None,t')::l) avoid env appc c'
+ share_names isgoal (n-1) ((Name id,None,t')::l) avoid env appc c'
(* If built with the f/n notation: we renounce to share names *)
| _ ->
if n>0 then warning "Detyping.detype: cannot factorize fix enough";
- let c = detype tenv avoid env c in
- let t = detype tenv avoid env t in
+ let c = detype isgoal avoid env c in
+ let t = detype isgoal avoid env t in
(List.rev l,c,t)
-and detype_eqn tenv avoid env constr construct_nargs branch =
+and detype_eqns isgoal avoid env ci computable constructs consnargsl bl =
+ try
+ if !Options.raw_print or not (reverse_matching ()) then raise Exit;
+ let mat = build_tree Anonymous isgoal (avoid,env) ci bl in
+ List.map (fun (pat,((avoid,env),c)) -> (dl,[],[pat],detype isgoal avoid env c))
+ mat
+ with _ ->
+ Array.to_list
+ (array_map3 (detype_eqn isgoal avoid env) constructs consnargsl bl)
+
+and detype_eqn isgoal avoid env constr construct_nargs branch =
let make_pat x avoid env b ids =
if force_wildcard () & noccurn 1 b then
- PatVar (dummy_loc,Anonymous),avoid,(add_name Anonymous env),ids
- else
+ PatVar (dl,Anonymous),avoid,(add_name Anonymous env),ids
+ else
let id = next_name_away_with_default "x" x avoid in
- PatVar (dummy_loc,Name id),id::avoid,(add_name (Name id) env),id::ids
+ PatVar (dl,Name id),id::avoid,(add_name (Name id) env),id::ids
in
let rec buildrec ids patlist avoid env n b =
if n=0 then
- (dummy_loc, ids,
- [PatCstr(dummy_loc, constr, List.rev patlist,Anonymous)],
- detype tenv avoid env b)
+ (dl, ids,
+ [PatCstr(dl, constr, List.rev patlist,Anonymous)],
+ detype isgoal avoid env b)
else
match kind_of_term b with
| Lambda (x,_,b) ->
@@ -465,7 +511,7 @@ and detype_eqn tenv avoid env constr construct_nargs branch =
let pat,new_avoid,new_env,new_ids = make_pat x avoid env b ids in
buildrec new_ids (pat::patlist) new_avoid new_env (n-1) b
- | Cast (c,_) -> (* Oui, il y a parfois des cast *)
+ | Cast (c,_,_) -> (* Oui, il y a parfois des cast *)
buildrec ids patlist avoid env n c
| _ -> (* eta-expansion : n'arrivera plus lorsque tous les
@@ -479,14 +525,123 @@ and detype_eqn tenv avoid env constr construct_nargs branch =
in
buildrec [] [] avoid env construct_nargs branch
-and detype_binder tenv bk avoid env na ty c =
+and detype_binder isgoal bk avoid env na ty c =
let na',avoid' =
if bk = BLetIn then
- concrete_let_name (fst tenv) avoid env na c
+ concrete_let_name isgoal avoid env na c
else
- concrete_name (fst tenv) avoid env na c in
- let r = detype tenv avoid' (add_name na' env) c in
+ concrete_name isgoal avoid env na c in
+ let r = detype isgoal avoid' (add_name na' env) c in
match bk with
- | BProd -> RProd (dummy_loc, na',detype tenv avoid env ty, r)
- | BLambda -> RLambda (dummy_loc, na',detype tenv avoid env ty, r)
- | BLetIn -> RLetIn (dummy_loc, na',detype tenv avoid env ty, r)
+ | BProd -> RProd (dl, na',detype isgoal avoid env ty, r)
+ | BLambda -> RLambda (dl, na',detype isgoal avoid env ty, r)
+ | BLetIn -> RLetIn (dl, na',detype isgoal avoid env ty, r)
+
+(**********************************************************************)
+(* Module substitution: relies on detyping *)
+
+let rec subst_cases_pattern subst pat =
+ match pat with
+ | PatVar _ -> pat
+ | PatCstr (loc,((kn,i),j),cpl,n) ->
+ let kn' = subst_kn subst kn
+ and cpl' = list_smartmap (subst_cases_pattern subst) cpl in
+ if kn' == kn && cpl' == cpl then pat else
+ PatCstr (loc,((kn',i),j),cpl',n)
+
+let rec subst_rawconstr subst raw =
+ match raw with
+ | RRef (loc,ref) ->
+ let ref',t = subst_global subst ref in
+ if ref' == ref then raw else
+ detype false [] [] t
+
+ | RVar _ -> raw
+ | REvar _ -> raw
+ | RPatVar _ -> raw
+
+ | RApp (loc,r,rl) ->
+ let r' = subst_rawconstr subst r
+ and rl' = list_smartmap (subst_rawconstr subst) rl in
+ if r' == r && rl' == rl then raw else
+ RApp(loc,r',rl')
+
+ | RLambda (loc,n,r1,r2) ->
+ let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ if r1' == r1 && r2' == r2 then raw else
+ RLambda (loc,n,r1',r2')
+
+ | RProd (loc,n,r1,r2) ->
+ let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ if r1' == r1 && r2' == r2 then raw else
+ RProd (loc,n,r1',r2')
+
+ | RLetIn (loc,n,r1,r2) ->
+ let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ if r1' == r1 && r2' == r2 then raw else
+ RLetIn (loc,n,r1',r2')
+
+ | RCases (loc,rtno,rl,branches) ->
+ let rtno' = option_smartmap (subst_rawconstr subst) rtno
+ and rl' = list_smartmap (fun (a,x as y) ->
+ let a' = subst_rawconstr subst a in
+ let (n,topt) = x in
+ let topt' = option_smartmap
+ (fun (loc,(sp,i),x as t) ->
+ let sp' = subst_kn subst sp in
+ if sp == sp' then t else (loc,(sp',i),x)) topt in
+ if a == a' && topt == topt' then y else (a',(n,topt'))) rl
+ and branches' = list_smartmap
+ (fun (loc,idl,cpl,r as branch) ->
+ let cpl' =
+ list_smartmap (subst_cases_pattern subst) cpl
+ and r' = subst_rawconstr subst r in
+ if cpl' == cpl && r' == r then branch else
+ (loc,idl,cpl',r'))
+ branches
+ in
+ if rtno' == rtno && rl' == rl && branches' == branches then raw else
+ RCases (loc,rtno',rl',branches')
+
+ | RLetTuple (loc,nal,(na,po),b,c) ->
+ let po' = option_smartmap (subst_rawconstr subst) po
+ and b' = subst_rawconstr subst b
+ and c' = subst_rawconstr subst c in
+ if po' == po && b' == b && c' == c then raw else
+ RLetTuple (loc,nal,(na,po'),b',c')
+
+ | RIf (loc,c,(na,po),b1,b2) ->
+ let po' = option_smartmap (subst_rawconstr subst) po
+ and b1' = subst_rawconstr subst b1
+ and b2' = subst_rawconstr subst b2
+ and c' = subst_rawconstr subst c in
+ if c' == c & po' == po && b1' == b1 && b2' == b2 then raw else
+ RIf (loc,c',(na,po'),b1',b2')
+
+ | RRec (loc,fix,ida,bl,ra1,ra2) ->
+ let ra1' = array_smartmap (subst_rawconstr subst) ra1
+ and ra2' = array_smartmap (subst_rawconstr subst) ra2 in
+ let bl' = array_smartmap
+ (list_smartmap (fun (na,obd,ty as dcl) ->
+ let ty' = subst_rawconstr subst ty in
+ let obd' = option_smartmap (subst_rawconstr subst) obd in
+ if ty'==ty & obd'==obd then dcl else (na,obd',ty')))
+ bl in
+ if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else
+ RRec (loc,fix,ida,bl',ra1',ra2')
+
+ | RSort _ -> raw
+
+ | RHole (loc,ImplicitArg (ref,i)) ->
+ let ref',_ = subst_global subst ref in
+ if ref' == ref then raw else
+ RHole (loc,InternalHole)
+ | RHole (loc, (BinderType _ | QuestionMark | CasesType |
+ InternalHole | TomatchTypeParameter _)) -> raw
+
+ | RCast (loc,r1,k,r2) ->
+ let r1' = subst_rawconstr subst r1 and r2' = subst_rawconstr subst r2 in
+ if r1' == r1 && r2' == r2 then raw else
+ RCast (loc,r1',k,r2')
+
+ | RDynamic _ -> raw
diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli
index c2a70928..0b35728c 100644
--- a/pretyping/detyping.mli
+++ b/pretyping/detyping.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: detyping.mli,v 1.13.2.2 2004/07/16 19:30:44 herbelin Exp $ i*)
+(*i $Id: detyping.mli 7881 2006-01-16 14:03:05Z herbelin $ i*)
(*i*)
open Util
@@ -16,20 +16,27 @@ open Sign
open Environ
open Rawterm
open Termops
+open Mod_subst
(*i*)
-(* [detype env avoid nenv c] turns [c], typed in [env], into a rawconstr. *)
-(* De Bruijn indexes are turned to bound names, avoiding names in [avoid] *)
+val subst_cases_pattern : substitution -> cases_pattern -> cases_pattern
-val detype : bool * env -> identifier list -> names_context -> constr ->
- rawconstr
+val subst_rawconstr : substitution -> rawconstr -> rawconstr
+
+(* [detype isgoal avoid ctx c] turns a closed [c], into a rawconstr *)
+(* de Bruijn indexes are turned to bound names, avoiding names in [avoid] *)
+(* [isgoal] tells if naming must avoid global-level synonyms as intro does *)
+(* [ctx] gives the names of the free variables *)
+
+val detype : bool -> identifier list -> names_context -> constr -> rawconstr
val detype_case :
bool -> ('a -> rawconstr) ->
- (constructor -> int -> 'a -> loc * identifier list * cases_pattern list *
- rawconstr) -> ('a -> int -> bool) ->
- env -> identifier list -> inductive -> case_style ->
- 'a option -> int -> 'a -> 'a array -> rawconstr
+ (constructor array -> int array -> 'a array ->
+ (loc * identifier list * cases_pattern list * rawconstr) list) ->
+ ('a -> int -> bool) ->
+ identifier list -> inductive * case_style * int * int array * int ->
+ 'a option -> 'a -> 'a array -> rawconstr
(* look for the index of a named var or a nondep var as it is renamed *)
val lookup_name_as_renamed : env -> constr -> identifier -> int option
diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml
index 2264f82b..2b04b693 100644
--- a/pretyping/evarconv.ml
+++ b/pretyping/evarconv.ml
@@ -6,20 +6,21 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evarconv.ml,v 1.44.6.3 2005/11/29 21:40:52 letouzey Exp $ *)
+(* $Id: evarconv.ml 8111 2006-03-02 17:23:41Z herbelin $ *)
open Util
open Names
open Term
+open Reduction
open Reductionops
open Closure
-open Instantiate
open Environ
open Typing
open Classops
open Recordops
open Evarutil
open Libnames
+open Evd
type flex_kind_of_term =
| Rigid of constr
@@ -69,7 +70,7 @@ let evar_apprec_nobeta env isevars stack c =
let (t,stack as s') = apprec_nobeta env (evars_of isevars) s in
match kind_of_term t with
| Evar (n,_ as ev) when Evd.is_defined (evars_of isevars) n ->
- aux (existential_value (evars_of isevars) ev, stack)
+ aux (Evd.existential_value (evars_of isevars) ev, stack)
| _ -> (t, list_of_stack stack)
in aux (c, append_stack (Array.of_list stack) empty_stack)
*)
@@ -77,10 +78,10 @@ let evar_apprec_nobeta env isevars stack c =
let evar_apprec env isevars stack c =
let sigma = evars_of isevars in
let rec aux s =
- let (t,stack as s') = Reductionops.apprec env sigma s in
+ let (t,stack) = Reductionops.apprec env sigma s in
match kind_of_term t with
| Evar (n,_ as ev) when Evd.is_defined sigma n ->
- aux (existential_value sigma ev, stack)
+ aux (Evd.existential_value sigma ev, stack)
| _ -> (t, list_of_stack stack)
in aux (c, append_stack (Array.of_list stack) empty_stack)
@@ -99,11 +100,11 @@ let apprec_nohdbeta env isevars c =
(t2 us2) = (cstr us)
extra_args1 = extra_args2
- by finding a record R and an object c := [xs:bs](Build_R a1..am v1..vn)
+ by finding a record R and an object c := [xs:bs](Build_R params v1..vn)
with vi = (cstr us), for which we know that the i-th projection proji
satisfies
- (proji params c) = (cstr us)
+ (proji params (c xs)) = (cstr us)
Rem: such objects, usable for conversion, are defined in the objdef
table; practically, it amounts to "canonically" equip t2 into a
@@ -112,10 +113,10 @@ let apprec_nohdbeta env isevars c =
let check_conv_record (t1,l1) (t2,l2) =
try
- let proji = reference_of_constr t1 in
- let cstr = reference_of_constr t2 in
+ let proji = global_of_constr t1 in
+ let cstr = global_of_constr t2 in
let { o_DEF = c; o_TABS = bs; o_TPARAMS = params; o_TCOMPS = us } =
- objdef_info (proji, cstr) in
+ lookup_canonical_conversion (proji, cstr) in
let params1, c1, extra_args1 =
match list_chop (List.length params) l1 with
| params1, c1::extra_args1 -> params1, c1, extra_args1
@@ -126,9 +127,47 @@ let check_conv_record (t1,l1) (t2,l2) =
raise Not_found
-(* Precondition: one of the terms of the pb is an uninstanciated evar,
+(* Precondition: one of the terms of the pb is an uninstantiated evar,
* possibly applied to arguments. *)
+let rec ise_try isevars = function
+ [] -> assert false
+ | [f] -> f isevars
+ | f1::l ->
+ let (isevars',b) = f1 isevars in
+ if b then (isevars',b) else ise_try isevars l
+
+let ise_and isevars l =
+ let rec ise_and i = function
+ [] -> assert false
+ | [f] -> f i
+ | f1::l ->
+ let (i',b) = f1 i in
+ if b then ise_and i' l else (isevars,false) in
+ ise_and isevars l
+
+let ise_list2 isevars f l1 l2 =
+ let rec ise_list2 i l1 l2 =
+ match l1,l2 with
+ [], [] -> (i, true)
+ | [x], [y] -> f i x y
+ | x::l1, y::l2 ->
+ let (i',b) = f i x y in
+ if b then ise_list2 i' l1 l2 else (isevars,false)
+ | _ -> (isevars, false) in
+ ise_list2 isevars l1 l2
+
+let ise_array2 isevars f v1 v2 =
+ let rec allrec i = function
+ | -1 -> (i,true)
+ | n ->
+ let (i',b) = f i v1.(n) v2.(n) in
+ if b then allrec i' (n-1) else (isevars,false)
+ in
+ let lv1 = Array.length v1 in
+ if lv1 = Array.length v2 then allrec isevars (pred lv1)
+ else (isevars,false)
+
let rec evar_conv_x env isevars pbty term1 term2 =
let sigma = evars_of isevars in
let term1 = whd_castappevar sigma term1 in
@@ -138,15 +177,15 @@ let rec evar_conv_x env isevars pbty term1 term2 =
true
else
*)
- (* Maybe convertible but since reducing can erase evars which [evar_apprec]*)
- (* could have found, we do it only if the terms are free of evar *)
- (not (has_undefined_isevars isevars term1) &
- not (has_undefined_isevars isevars term2) &
- is_fconv pbty env (evars_of isevars) term1 term2)
- or
- if ise_undefined isevars term1 then
+ (* Maybe convertible but since reducing can erase evars which [evar_apprec]
+ could have found, we do it only if the terms are free of evar.
+ Note: incomplete heuristic... *)
+ if is_ground_term isevars term1 && is_ground_term isevars term2 &
+ is_fconv pbty env (evars_of isevars) term1 term2 then
+ (isevars,true)
+ else if is_undefined_evar isevars term1 then
solve_simple_eqn evar_conv_x env isevars (pbty,destEvar term1,term2)
- else if ise_undefined isevars term2 then
+ else if is_undefined_evar isevars term2 then
solve_simple_eqn evar_conv_x env isevars (pbty,destEvar term2,term1)
else
let (t1,l1) = apprec_nohdbeta env isevars term1 in
@@ -154,7 +193,7 @@ let rec evar_conv_x env isevars pbty term1 term2 =
if (head_is_embedded_evar isevars t1 & not(is_eliminator t2))
or (head_is_embedded_evar isevars t2 & not(is_eliminator t1))
then
- (add_conv_pb isevars (pbty,applist(t1,l1),applist(t2,l2)); true)
+ (add_conv_pb (pbty,applist(t1,l1),applist(t2,l2)) isevars, true)
else
evar_eqappr_x env isevars pbty (t1,l1) (t2,l2)
@@ -162,67 +201,81 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) =
(* Evar must be undefined since we have whd_ised *)
match (flex_kind_of_term term1 l1, flex_kind_of_term term2 l2) with
| Flexible (sp1,al1 as ev1), Flexible (sp2,al2 as ev2) ->
- let f1 () =
+ let f1 i =
if List.length l1 > List.length l2 then
let (deb1,rest1) = list_chop (List.length l1-List.length l2) l1 in
- solve_simple_eqn evar_conv_x env isevars
- (pbty,ev2,applist(term1,deb1))
- & list_for_all2eq (evar_conv_x env isevars CONV) rest1 l2
+ ise_and i
+ [(fun i -> solve_simple_eqn evar_conv_x env i
+ (pbty,ev2,applist(term1,deb1)));
+ (fun i -> ise_list2 i
+ (fun i -> evar_conv_x env i CONV) rest1 l2)]
else
let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in
- solve_simple_eqn evar_conv_x env isevars
- (pbty,ev1,applist(term2,deb2))
- & list_for_all2eq (evar_conv_x env isevars CONV) l1 rest2
- and f2 () =
- (sp1 = sp2)
- & (array_for_all2 (evar_conv_x env isevars CONV) al1 al2)
- & (list_for_all2eq (evar_conv_x env isevars CONV) l1 l2)
+ ise_and i
+ [(fun i -> solve_simple_eqn evar_conv_x env i
+ (pbty,ev1,applist(term2,deb2)));
+ (fun i -> ise_list2 i
+ (fun i -> evar_conv_x env i CONV) l1 rest2)]
+ and f2 i =
+ if sp1 = sp2 then
+ ise_and i
+ [(fun i -> ise_array2 i
+ (fun i -> evar_conv_x env i CONV) al1 al2);
+ (fun i -> ise_list2 i
+ (fun i -> evar_conv_x env i CONV) l1 l2)]
+ else (i,false)
in
ise_try isevars [f1; f2]
| Flexible ev1, MaybeFlexible flex2 ->
- let f1 () =
- (List.length l1 <= List.length l2) &
- let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in
- (* First compare extra args for better failure message *)
- list_for_all2eq (evar_conv_x env isevars CONV) l1 rest2 &
- evar_conv_x env isevars pbty term1 (applist(term2,deb2))
- and f4 () =
+ let f1 i =
+ if List.length l1 <= List.length l2 then
+ let (deb2,rest2) =
+ list_chop (List.length l2-List.length l1) l2 in
+ ise_and i
+ (* First compare extra args for better failure message *)
+ [(fun i -> ise_list2 i
+ (fun i -> evar_conv_x env i CONV) l1 rest2);
+ (fun i -> evar_conv_x env i pbty term1 (applist(term2,deb2)))]
+ else (i,false)
+ and f4 i =
match eval_flexible_term env flex2 with
| Some v2 ->
- evar_eqappr_x env isevars pbty
- appr1 (evar_apprec env isevars l2 v2)
- | None -> false
+ evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
+ | None -> (i,false)
in
ise_try isevars [f1; f4]
| MaybeFlexible flex1, Flexible ev2 ->
- let f1 () =
- (List.length l2 <= List.length l1) &
- let (deb1,rest1) = list_chop (List.length l1-List.length l2) l1 in
- (* First compare extra args for better failure message *)
- list_for_all2eq (evar_conv_x env isevars CONV) rest1 l2 &
- evar_conv_x env isevars pbty (applist(term1,deb1)) term2
- and f4 () =
+ let f1 i =
+ if List.length l2 <= List.length l1 then
+ let (deb1,rest1) = list_chop (List.length l1-List.length l2) l1 in
+ ise_and i
+ (* First compare extra args for better failure message *)
+ [(fun i -> ise_list2 i
+ (fun i -> evar_conv_x env i CONV) rest1 l2);
+ (fun i -> evar_conv_x env i pbty (applist(term1,deb1)) term2)]
+ else (i,false)
+ and f4 i =
match eval_flexible_term env flex1 with
| Some v1 ->
- evar_eqappr_x env isevars pbty
- (evar_apprec env isevars l1 v1) appr2
- | None -> false
+ evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
+ | None -> (i,false)
in
ise_try isevars [f1; f4]
| MaybeFlexible flex1, MaybeFlexible flex2 ->
- let f2 () =
- (flex1 = flex2)
- & (List.length l1 = List.length l2)
- & (list_for_all2eq (evar_conv_x env isevars CONV) l1 l2)
- and f3 () =
- (try conv_record env isevars
+ let f2 i =
+ if flex1 = flex2 then
+ ise_list2 i (fun i -> evar_conv_x env i CONV) l1 l2
+ else (i,false)
+ and f3 i =
+ (try conv_record env i
(try check_conv_record appr1 appr2
with Not_found -> check_conv_record appr2 appr1)
- with _ -> false)
- and f4 () =
+(* TODO: remove this _ !!! *)
+ with _ -> (i,false))
+ and f4 i =
(* heuristic: unfold second argument first, exception made
if the first argument is a beta-redex (expand a constant
only if necessary) *)
@@ -232,87 +285,98 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) =
| _ -> eval_flexible_term env flex2 in
match val2 with
| Some v2 ->
- evar_eqappr_x env isevars pbty
- appr1 (evar_apprec env isevars l2 v2)
+ evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
| None ->
match eval_flexible_term env flex1 with
| Some v1 ->
- evar_eqappr_x env isevars pbty
- (evar_apprec env isevars l1 v1) appr2
- | None -> false
+ evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
+ | None -> (i,false)
in
ise_try isevars [f2; f3; f4]
| Flexible ev1, Rigid _ ->
- (List.length l1 <= List.length l2) &
- let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in
- (* First compare extra args for better failure message *)
- list_for_all2eq (evar_conv_x env isevars CONV) l1 rest2 &
- solve_simple_eqn evar_conv_x env isevars
- (pbty,ev1,applist(term2,deb2))
-
+ if (List.length l1 <= List.length l2) then
+ let (deb2,rest2) = list_chop (List.length l2-List.length l1) l2 in
+ ise_and isevars
+ (* First compare extra args for better failure message *)
+ [(fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) l1 rest2);
+ (fun i ->
+ (* Then instantiate evar unless already done by unifying args *)
+ let t2 = applist(term2,deb2) in
+ if is_defined_evar i ev1 then
+ evar_conv_x env i pbty (mkEvar ev1) t2
+ else
+ solve_simple_eqn evar_conv_x env i (pbty,ev1,t2))]
+ else (isevars,false)
| Rigid _, Flexible ev2 ->
- (List.length l2 <= List.length l1) &
- let (deb1,rest1) = list_chop (List.length l1-List.length l2) l1 in
- (* First compare extra args for better failure message *)
- list_for_all2eq (evar_conv_x env isevars CONV) rest1 l2 &
- solve_simple_eqn evar_conv_x env isevars
- (pbty,ev2,applist(term1,deb1))
-
+ if List.length l2 <= List.length l1 then
+ let (deb1,rest1) = list_chop (List.length l1-List.length l2) l1 in
+ ise_and isevars
+ (* First compare extra args for better failure message *)
+ [(fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) rest1 l2);
+ (fun i ->
+ (* Then instantiate evar unless already done by unifying args *)
+ let t1 = applist(term1,deb1) in
+ if is_defined_evar i ev2 then
+ evar_conv_x env i pbty t1 (mkEvar ev2)
+ else
+ solve_simple_eqn evar_conv_x env i (pbty,ev2,t1))]
+ else (isevars,false)
| MaybeFlexible flex1, Rigid _ ->
- let f3 () =
- (try conv_record env isevars (check_conv_record appr1 appr2)
- with _ -> false)
- and f4 () =
+ let f3 i =
+ (try conv_record env i (check_conv_record appr1 appr2)
+ with _ -> (i,false))
+ and f4 i =
match eval_flexible_term env flex1 with
| Some v1 ->
- evar_eqappr_x env isevars pbty
- (evar_apprec env isevars l1 v1) appr2
- | None -> false
+ evar_eqappr_x env i pbty (evar_apprec env i l1 v1) appr2
+ | None -> (i,false)
in
ise_try isevars [f3; f4]
| Rigid _ , MaybeFlexible flex2 ->
- let f3 () =
- (try (conv_record env isevars (check_conv_record appr2 appr1))
- with _ -> false)
- and f4 () =
+ let f3 i =
+ (try (conv_record env i (check_conv_record appr2 appr1))
+ with _ -> (i,false))
+ and f4 i =
match eval_flexible_term env flex2 with
| Some v2 ->
- evar_eqappr_x env isevars pbty
- appr1 (evar_apprec env isevars l2 v2)
- | None -> false
+ evar_eqappr_x env i pbty appr1 (evar_apprec env i l2 v2)
+ | None -> (i,false)
in
ise_try isevars [f3; f4]
| Rigid c1, Rigid c2 -> match kind_of_term c1, kind_of_term c2 with
- | Cast (c1,_), _ -> evar_eqappr_x env isevars pbty (c1,l1) appr2
+ | Cast (c1,_,_), _ -> evar_eqappr_x env isevars pbty (c1,l1) appr2
- | _, Cast (c2,_) -> evar_eqappr_x env isevars pbty appr1 (c2,l2)
+ | _, Cast (c2,_,_) -> evar_eqappr_x env isevars pbty appr1 (c2,l2)
- | Sort s1, Sort s2 when l1=[] & l2=[] -> base_sort_cmp pbty s1 s2
+ | Sort s1, Sort s2 when l1=[] & l2=[] ->
+ (isevars,base_sort_cmp pbty s1 s2)
| Lambda (na,c1,c'1), Lambda (_,c2,c'2) when l1=[] & l2=[] ->
- evar_conv_x env isevars CONV c1 c2
- &
- (let c = nf_evar (evars_of isevars) c1 in
- evar_conv_x (push_rel (na,None,c) env) isevars CONV c'1 c'2)
+ ise_and isevars
+ [(fun i -> evar_conv_x env i CONV c1 c2);
+ (fun i ->
+ let c = nf_evar (evars_of i) c1 in
+ evar_conv_x (push_rel (na,None,c) env) i CONV c'1 c'2)]
| LetIn (na,b1,t1,c'1), LetIn (_,b2,_,c'2) ->
- let f1 () =
- evar_conv_x env isevars CONV b1 b2
- &
- (let b = nf_evar (evars_of isevars) b1 in
- let t = nf_evar (evars_of isevars) t1 in
- evar_conv_x (push_rel (na,Some b,t) env) isevars pbty c'1 c'2)
- & (List.length l1 = List.length l2)
- & (List.for_all2 (evar_conv_x env isevars CONV) l1 l2)
- and f2 () =
- let appr1 = evar_apprec env isevars l1 (subst1 b1 c'1)
- and appr2 = evar_apprec env isevars l2 (subst1 b2 c'2)
- in evar_eqappr_x env isevars pbty appr1 appr2
+ let f1 i =
+ ise_and i
+ [(fun i -> evar_conv_x env i CONV b1 b2);
+ (fun i ->
+ let b = nf_evar (evars_of i) b1 in
+ let t = nf_evar (evars_of i) t1 in
+ evar_conv_x (push_rel (na,Some b,t) env) i pbty c'1 c'2);
+ (fun i -> ise_list2 i
+ (fun i -> evar_conv_x env i CONV) l1 l2)]
+ and f2 i =
+ let appr1 = evar_apprec env i l1 (subst1 b1 c'1)
+ and appr2 = evar_apprec env i l2 (subst1 b2 c'2)
+ in evar_eqappr_x env i pbty appr1 appr2
in
ise_try isevars [f1; f2]
@@ -325,71 +389,102 @@ and evar_eqappr_x env isevars pbty (term1,l1 as appr1) (term2,l2 as appr2) =
in evar_eqappr_x env isevars pbty appr1 appr2
| Prod (n,c1,c'1), Prod (_,c2,c'2) when l1=[] & l2=[] ->
- evar_conv_x env isevars CONV c1 c2
- &
- (let c = nf_evar (evars_of isevars) c1 in
- evar_conv_x (push_rel (n,None,c) env) isevars pbty c'1 c'2)
+ ise_and isevars
+ [(fun i -> evar_conv_x env i CONV c1 c2);
+ (fun i ->
+ let c = nf_evar (evars_of i) c1 in
+ evar_conv_x (push_rel (n,None,c) env) i pbty c'1 c'2)]
| Ind sp1, Ind sp2 ->
- sp1=sp2
- & list_for_all2eq (evar_conv_x env isevars CONV) l1 l2
-
+ if sp1=sp2 then
+ ise_list2 isevars (fun i -> evar_conv_x env i CONV) l1 l2
+ else (isevars, false)
+
| Construct sp1, Construct sp2 ->
- sp1=sp2
- & list_for_all2eq (evar_conv_x env isevars CONV) l1 l2
+ if sp1=sp2 then
+ ise_list2 isevars (fun i -> evar_conv_x env i CONV) l1 l2
+ else (isevars, false)
| Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) ->
- evar_conv_x env isevars CONV p1 p2
- & evar_conv_x env isevars CONV c1 c2
- & (array_for_all2 (evar_conv_x env isevars CONV) cl1 cl2)
- & (list_for_all2eq (evar_conv_x env isevars CONV) l1 l2)
+ ise_and isevars
+ [(fun i -> evar_conv_x env i CONV p1 p2);
+ (fun i -> evar_conv_x env i CONV c1 c2);
+ (fun i -> ise_array2 i
+ (fun i -> evar_conv_x env i CONV) cl1 cl2);
+ (fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) l1 l2)]
| Fix (li1,(_,tys1,bds1 as recdef1)), Fix (li2,(_,tys2,bds2)) ->
- li1=li2
- & (array_for_all2 (evar_conv_x env isevars CONV) tys1 tys2)
- & (array_for_all2
- (evar_conv_x (push_rec_types recdef1 env) isevars CONV)
- bds1 bds2)
- & (list_for_all2eq (evar_conv_x env isevars CONV) l1 l2)
-
+ if li1=li2 then
+ ise_and isevars
+ [(fun i -> ise_array2 i
+ (fun i -> evar_conv_x env i CONV) tys1 tys2);
+ (fun i -> ise_array2 i
+ (fun i -> evar_conv_x (push_rec_types recdef1 env) i CONV)
+ bds1 bds2);
+ (fun i -> ise_list2 i
+ (fun i -> evar_conv_x env i CONV) l1 l2)]
+ else (isevars,false)
| CoFix (i1,(_,tys1,bds1 as recdef1)), CoFix (i2,(_,tys2,bds2)) ->
- i1=i2
- & (array_for_all2 (evar_conv_x env isevars CONV) tys1 tys2)
- & (array_for_all2
- (evar_conv_x (push_rec_types recdef1 env) isevars CONV)
- bds1 bds2)
- & (list_for_all2eq (evar_conv_x env isevars CONV) l1 l2)
-
- | (Meta _ | Lambda _), _ -> false
- | _, (Meta _ | Lambda _) -> false
-
- | (Ind _ | Construct _ | Sort _ | Prod _), _ -> false
- | _, (Ind _ | Construct _ | Sort _ | Prod _) -> false
+ if i1=i2 then
+ ise_and isevars
+ [(fun i -> ise_array2 i
+ (fun i -> evar_conv_x env i CONV) tys1 tys2);
+ (fun i -> ise_array2 i
+ (fun i -> evar_conv_x (push_rec_types recdef1 env) i CONV)
+ bds1 bds2);
+ (fun i -> ise_list2 i
+ (fun i -> evar_conv_x env i CONV) l1 l2)]
+ else (isevars,false)
+
+ | (Meta _ | Lambda _), _ -> (isevars,false)
+ | _, (Meta _ | Lambda _) -> (isevars,false)
+
+ | (Ind _ | Construct _ | Sort _ | Prod _), _ -> (isevars,false)
+ | _, (Ind _ | Construct _ | Sort _ | Prod _) -> (isevars,false)
| (App _ | Case _ | Fix _ | CoFix _),
- (App _ | Case _ | Fix _ | CoFix _) -> false
+ (App _ | Case _ | Fix _ | CoFix _) -> (isevars,false)
| (Rel _ | Var _ | Const _ | Evar _), _ -> assert false
| _, (Rel _ | Var _ | Const _ | Evar _) -> assert false
and conv_record env isevars (c,bs,(params,params1),(us,us2),(ts,ts1),c1) =
- let ks =
+ let (isevars',ks) =
List.fold_left
- (fun ks b ->
- let dloc = (dummy_loc,Rawterm.InternalHole) in
- (new_isevar isevars env dloc (substl ks b)) :: ks)
- [] bs
+ (fun (i,ks) b ->
+ let dloc = (dummy_loc,InternalHole) in
+ let (i',ev) = new_evar i env ~src:dloc (substl ks b) in
+ (i', ev :: ks))
+ (isevars,[]) bs
in
- (list_for_all2eq
- (fun u1 u -> evar_conv_x env isevars CONV u1 (substl ks u))
- us2 us)
- &
- (list_for_all2eq
- (fun x1 x -> evar_conv_x env isevars CONV x1 (substl ks x))
- params1 params)
- & (list_for_all2eq (evar_conv_x env isevars CONV) ts ts1)
- & (evar_conv_x env isevars CONV c1 (applist (c,(List.rev ks))))
+ ise_and isevars'
+ [(fun i ->
+ ise_list2 i
+ (fun i u1 u -> evar_conv_x env i CONV u1 (substl ks u))
+ us2 us);
+ (fun i ->
+ ise_list2 i
+ (fun i x1 x -> evar_conv_x env i CONV x1 (substl ks x))
+ params1 params);
+ (fun i -> ise_list2 i (fun i -> evar_conv_x env i CONV) ts ts1);
+ (fun i -> evar_conv_x env i CONV c1 (applist (c,(List.rev ks))))]
-let the_conv_x env isevars t1 t2 = evar_conv_x env isevars CONV t1 t2
-let the_conv_x_leq env isevars t1 t2 = evar_conv_x env isevars CUMUL t1 t2
+let the_conv_x env t1 t2 isevars =
+ match evar_conv_x env isevars CONV t1 t2 with
+ (evd',true) -> evd'
+ | _ -> raise Reduction.NotConvertible
+
+let the_conv_x_leq env t1 t2 isevars =
+ match evar_conv_x env isevars CUMUL t1 t2 with
+ (evd', true) -> evd'
+ | _ -> raise Reduction.NotConvertible
+let e_conv env isevars t1 t2 =
+ match evar_conv_x env !isevars CONV t1 t2 with
+ (evd',true) -> isevars := evd'; true
+ | _ -> false
+
+let e_cumul env isevars t1 t2 =
+ match evar_conv_x env !isevars CUMUL t1 t2 with
+ (evd',true) -> isevars := evd'; true
+ | _ -> false
diff --git a/pretyping/evarconv.mli b/pretyping/evarconv.mli
index 8785d855..a6f5b489 100644
--- a/pretyping/evarconv.mli
+++ b/pretyping/evarconv.mli
@@ -6,23 +6,30 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evarconv.mli,v 1.11.14.1 2004/07/16 19:30:44 herbelin Exp $ i*)
+(*i $Id: evarconv.mli 6109 2004-09-15 16:50:56Z barras $ i*)
(*i*)
open Term
open Sign
open Environ
open Reductionops
-open Evarutil
+open Evd
(*i*)
-val the_conv_x : env -> evar_defs -> constr -> constr -> bool
+(* returns exception Reduction.NotConvertible if not unifiable *)
+val the_conv_x : env -> constr -> constr -> evar_defs -> evar_defs
+val the_conv_x_leq : env -> constr -> constr -> evar_defs -> evar_defs
-val the_conv_x_leq : env -> evar_defs -> constr -> constr -> bool
+(* The same function resolving evars by side-effect and
+ catching the exception *)
+val e_conv : env -> evar_defs ref -> constr -> constr -> bool
+val e_cumul : env -> evar_defs ref -> constr -> constr -> bool
(*i For debugging *)
-val evar_conv_x : env -> evar_defs -> conv_pb -> constr -> constr -> bool
+val evar_conv_x :
+ env -> evar_defs -> conv_pb -> constr -> constr -> evar_defs * bool
val evar_eqappr_x :
env -> evar_defs ->
- conv_pb -> constr * constr list -> constr * constr list -> bool
+ conv_pb -> constr * constr list -> constr * constr list ->
+ evar_defs * bool
(*i*)
diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml
index 4337c0fc..aeaaefef 100644
--- a/pretyping/evarutil.ml
+++ b/pretyping/evarutil.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evarutil.ml,v 1.64.2.5 2004/12/09 14:45:38 herbelin Exp $ *)
+(* $Id: evarutil.ml 8695 2006-04-10 16:33:52Z msozeau $ *)
open Util
open Pp
@@ -18,9 +18,7 @@ open Termops
open Sign
open Environ
open Evd
-open Instantiate
open Reductionops
-open Indrec
open Pretype_errors
@@ -50,7 +48,7 @@ let whd_castappevar_stack sigma c =
match kind_of_term c with
| Evar (ev,args) when Evd.in_dom sigma ev & Evd.is_defined sigma ev ->
whrec (existential_value sigma (ev,args), l)
- | Cast (c,_) -> whrec (c, l)
+ | Cast (c,_,_) -> whrec (c, l)
| App (f,args) -> whrec (f, Array.fold_right (fun a l -> a::l) args l)
| _ -> s
in
@@ -64,207 +62,121 @@ let jl_nf_evar = Pretype_errors.jl_nf_evar
let jv_nf_evar = Pretype_errors.jv_nf_evar
let tj_nf_evar = Pretype_errors.tj_nf_evar
-(**********************)
-(* Creating new evars *)
-(**********************)
-
-let evar_env evd = Global.env_of_context evd.evar_hyps
-
-(* Generator of existential names *)
-let new_evar =
- let evar_ctr = ref 0 in
- fun () -> incr evar_ctr; existential_of_int !evar_ctr
-
-let make_evar_instance env =
- fold_named_context
- (fun env (id, b, _) l -> (*if b=None then*) mkVar id :: l (*else l*))
- env ~init:[]
+let nf_evar_info evc info =
+ { evar_concl = Reductionops.nf_evar evc info.evar_concl;
+ evar_hyps = map_named_val (Reductionops.nf_evar evc) info.evar_hyps;
+ evar_body = info.evar_body}
-(* create an untyped existential variable *)
-let new_evar_in_sign env =
- let ev = new_evar () in
- mkEvar (ev, Array.of_list (make_evar_instance env))
+let nf_evars evm = Evd.fold (fun ev evi evm' -> Evd.add evm' ev (nf_evar_info evm evi))
+ evm Evd.empty
-(*------------------------------------*
- * functional operations on evar sets *
- *------------------------------------*)
+let nf_evar_defs isevars = Evd.evars_reset_evd (nf_evars (Evd.evars_of isevars)) isevars
-(* All ids of sign must be distincts! *)
-let new_isevar_sign env sigma typ instance =
- let sign = named_context env in
- if not (list_distinct (ids_of_named_context sign)) then
- error "new_isevar_sign: two vars have the same name";
- let newev = new_evar() in
- let info = { evar_concl = typ; evar_hyps = sign;
- evar_body = Evar_empty } in
- (Evd.add sigma newev info, mkEvar (newev,Array.of_list instance))
+let nf_isevar isevars = nf_evar (Evd.evars_of isevars)
+let j_nf_isevar isevars = j_nf_evar (Evd.evars_of isevars)
+let jl_nf_isevar isevars = jl_nf_evar (Evd.evars_of isevars)
+let jv_nf_isevar isevars = jv_nf_evar (Evd.evars_of isevars)
+let tj_nf_isevar isevars = tj_nf_evar (Evd.evars_of isevars)
-(* We don't try to guess in which sort the type should be defined, since
- any type has type Type. May cause some trouble, but not so far... *)
-let new_Type () = mkType (new_univ ())
-
-let new_Type_sort () = Type (new_univ ())
+(**********************)
+(* Creating new metas *)
+(**********************)
-let judge_of_new_Type () = Typeops.judge_of_type (new_univ ())
-(*
-let new_Type () = mkType dummy_univ
-
-let new_Type_sort () = Type dummy_univ
-
-let judge_of_new_Type () =
- { uj_val = mkSort (Type dummy_univ);
- uj_type = mkSort (Type dummy_univ) }
-*)
-
-(* This refreshes universes in types; works only for inferred types (i.e. for
- types of the form (x1:A1)...(xn:An)B with B a sort or an atom in
- head normal form) *)
-let refresh_universes t =
- let modified = ref false in
- let rec refresh t = match kind_of_term t with
- | Sort (Type _) -> modified := true; new_Type ()
- | Prod (na,u,v) -> mkProd (na,u,refresh v)
- | _ -> t in
- let t' = refresh t in
- if !modified then t' else t
+(* Generator of metavariables *)
+let new_meta =
+ let meta_ctr = ref 0 in
+ fun () -> incr meta_ctr; !meta_ctr
-(* Declaring any type to be in the sort Type shouldn't be harmful since
- cumulativity now includes Prop and Set in Type. *)
-let new_type_var env sigma =
- let instance = make_evar_instance env in
- new_isevar_sign env sigma (new_Type ()) instance
-
-let split_evar_to_arrow sigma (ev,args) =
- let evd = Evd.map sigma ev in
- let evenv = evar_env evd in
- let (sigma1,dom) = new_type_var evenv sigma in
- let hyps = evd.evar_hyps in
- let nvar = next_ident_away (id_of_string "x") (ids_of_named_context hyps) in
- let newenv = push_named (nvar, None, dom) evenv in
- let (sigma2,rng) = new_type_var newenv sigma1 in
- let x = named_hd newenv dom Anonymous in
- let prod = mkProd (x, dom, subst_var nvar rng) in
- let sigma3 = Evd.define sigma2 ev prod in
- let evdom = fst (destEvar dom), args in
- let evrng =
- fst (destEvar rng), array_cons (mkRel 1) (Array.map (lift 1) args) in
- let prod' = mkProd (x, mkEvar evdom, mkEvar evrng) in
- (sigma3,prod', evdom, evrng)
+let mk_new_meta () = mkMeta(new_meta())
-(* Redefines an evar with a smaller context (i.e. it may depend on less
- * variables) such that c becomes closed.
- * Example: in [x:?1; y:(list ?2)] <?3>x=y /\ x=(nil bool)
- * ?3 <-- ?1 no pb: env of ?3 is larger than ?1's
- * ?1 <-- (list ?2) pb: ?2 may depend on x, but not ?1.
- * What we do is that ?2 is defined by a new evar ?4 whose context will be
- * a prefix of ?2's env, included in ?1's env. *)
+let collect_evars emap c =
+ let rec collrec acc c =
+ match kind_of_term c with
+ | Evar (k,_) ->
+ if Evd.in_dom emap k & not (Evd.is_defined emap k) then k::acc
+ else (* No recursion on the evar instantiation *) acc
+ | _ ->
+ fold_constr collrec acc c in
+ list_uniquize (collrec [] c)
+
+let push_dependent_evars sigma emap =
+ Evd.fold (fun ev {evar_concl = ccl} (sigma',emap') ->
+ List.fold_left
+ (fun (sigma',emap') ev ->
+ (Evd.add sigma' ev (Evd.map emap' ev),Evd.rmv emap' ev))
+ (sigma',emap') (collect_evars emap' ccl))
+ emap (sigma,emap)
+
+(* replaces a mapping of existentials into a mapping of metas.
+ Problem if an evar appears in the type of another one (pops anomaly) *)
+let evars_to_metas sigma (emap, c) =
+ let sigma',emap' = push_dependent_evars sigma emap in
+ let change_exist evar =
+ let ty = nf_betaiota (nf_evar emap (existential_type emap evar)) in
+ let n = new_meta() in
+ mkCast (mkMeta n, DEFAULTcast, ty) in
+ let rec replace c =
+ match kind_of_term c with
+ Evar (k,_ as ev) when Evd.in_dom emap' k -> change_exist ev
+ | _ -> map_constr replace c in
+ (sigma', replace c)
+
+(* The list of non-instantiated existential declarations *)
+
+let non_instantiated sigma =
+ let listev = to_list sigma in
+ List.fold_left
+ (fun l (ev,evd) ->
+ if evd.evar_body = Evar_empty then
+ ((ev,nf_evar_info sigma evd)::l) else l)
+ [] listev
+
+(*************************************)
+(* Metas *)
+
+let meta_value evd mv =
+ let rec valrec mv =
+ try
+ let b = meta_fvalue evd mv in
+ instance
+ (List.map (fun mv' -> (mv',valrec mv')) (Metaset.elements b.freemetas))
+ b.rebus
+ with Anomaly _ | Not_found ->
+ mkMeta mv
+ in
+ valrec mv
-let do_restrict_hyps sigma ev args =
- let args = Array.to_list args in
- let evd = Evd.map sigma ev in
- let env = evar_env evd in
- let hyps = evd.evar_hyps in
- let (sign,ncargs) = list_filter2 (fun _ a -> closed0 a) (hyps,args) in
- let env' = reset_with_named_context sign env in
- let (sigma',nc) = new_isevar_sign env' sigma evd.evar_concl ncargs in
- let nc = refresh_universes nc in (* needed only if nc is an inferred type *)
- let sigma'' = Evd.define sigma' ev nc in
- (sigma'', nc)
+let meta_instance env b =
+ let c_sigma =
+ List.map
+ (fun mv -> (mv,meta_value env mv)) (Metaset.elements b.freemetas)
+ in
+ instance c_sigma b.rebus
+let nf_meta env c = meta_instance env (mk_freelisted c)
+(**********************)
+(* Creating new evars *)
+(**********************)
+(* Generator of existential names *)
+let new_untyped_evar =
+ let evar_ctr = ref 0 in
+ fun () -> incr evar_ctr; existential_of_int !evar_ctr
(*------------------------------------*
- * operations on the evar constraints *
+ * functional operations on evar sets *
*------------------------------------*)
-type evar_constraint = conv_pb * constr * constr
-type evar_defs =
- { mutable evars : Evd.evar_map;
- mutable conv_pbs : evar_constraint list;
- mutable history : (existential_key * (loc * Rawterm.hole_kind)) list }
-
-let create_evar_defs evd = { evars=evd; conv_pbs=[]; history=[] }
-let evars_of d = d.evars
-let evars_reset_evd evd d = d.evars <- evd
-let add_conv_pb d pb = d.conv_pbs <- pb::d.conv_pbs
-let evar_source ev d =
- try List.assoc ev d.history
- with Failure _ -> (dummy_loc, Rawterm.InternalHole)
-
-(* ise_try [f1;...;fn] tries fi() for i=1..n, restoring the evar constraints
- * when fi returns false or an exception. Returns true if one of the fi
- * returns true, and false if every fi return false (in the latter case,
- * the evar constraints are restored).
- *)
-let ise_try isevars l =
- let u = isevars.evars in
- let rec test = function
- [] -> isevars.evars <- u; false
- | f::l ->
- (try f() with reraise -> isevars.evars <- u; raise reraise)
- or (isevars.evars <- u; test l)
- in test l
-
-
-
-(* say if the section path sp corresponds to an existential *)
-let ise_in_dom isevars sp = Evd.in_dom isevars.evars sp
-
-(* map the given section path to the enamed_declaration *)
-let ise_map isevars sp = Evd.map isevars.evars sp
-
-(* define the existential of section path sp as the constr body *)
-let ise_define isevars sp body =
- let body = refresh_universes body in (* needed only if an inferred type *)
- isevars.evars <- Evd.define isevars.evars sp body
-
-let is_defined_evar isevars (n,_) = Evd.is_defined isevars.evars n
-
-(* Does k corresponds to an (un)defined existential ? *)
-let ise_undefined isevars c = match kind_of_term c with
- | Evar ev -> not (is_defined_evar isevars ev)
- | _ -> false
-
-let need_restriction isevars args = not (array_for_all closed0 args)
-
-
-(* We try to instanciate the evar assuming the body won't depend
- * on arguments that are not Rels or Vars, or appearing several times.
- *)
-(* Note: error_not_clean should not be an error: it simply means that the
- * conversion test that lead to the faulty call to [real_clean] should return
- * false. The problem is that we won't get the right error message.
- *)
-
-let real_clean env isevars ev args rhs =
- let subst = List.map (fun (x,y) -> (y,mkVar x)) (filter_unique args) in
- let rec subs k t =
- match kind_of_term t with
- | Rel i ->
- if i<=k then t
- else (try List.assoc (mkRel (i-k)) subst with Not_found -> t)
- | Evar (ev,args) ->
- let args' = Array.map (subs k) args in
- if need_restriction isevars args' then
- if Evd.is_defined isevars.evars ev then
- subs k (existential_value isevars.evars (ev,args'))
- else begin
- let (sigma,rc) = do_restrict_hyps isevars.evars ev args' in
- isevars.evars <- sigma;
- isevars.history <-
- (fst (destEvar rc),evar_source ev isevars)::isevars.history;
- rc
- end
- else
- mkEvar (ev,args')
- | Var _ -> (try List.assoc t subst with Not_found -> t)
- | _ -> map_constr_with_binders succ subs k t
- in
- let body = subs 0 rhs in
- if not (closed0 body)
- then error_not_clean env isevars.evars ev body (evar_source ev isevars);
- body
+(* All ids of sign must be distincts! *)
+let new_evar_instance sign evd typ ?(src=(dummy_loc,InternalHole)) instance =
+ let ctxt = named_context_of_val sign in
+ assert (List.length instance = named_context_length ctxt);
+ if not (list_distinct (ids_of_named_context ctxt)) then
+ anomaly "new_evar_instance: two vars have the same name";
+ let newev = new_untyped_evar() in
+ (evar_declare sign newev typ ~src:src evd,
+ mkEvar (newev,Array.of_list instance))
let make_evar_instance_with_rel env =
let n = rel_context_length (rel_context env) in
@@ -279,7 +191,7 @@ let make_evar_instance_with_rel env =
let make_subst env args =
snd (fold_named_context
- (fun env (id,b,c) (args,l as g) ->
+ (fun env (id,b,c) (args,l) ->
match b, args with
| (* None *) _ , a::rest -> (rest, (id,a)::l)
(* | Some _, _ -> g*)
@@ -290,28 +202,175 @@ let make_subst env args =
(* Converting the env into the sign of the evar to define *)
let push_rel_context_to_named_context env =
- let sign0 = named_context env in
- let (subst,_,sign) =
+ let (subst,_,env) =
Sign.fold_rel_context
- (fun (na,c,t) (subst,avoid,sign) ->
+ (fun (na,c,t) (subst,avoid,env) ->
let na = if na = Anonymous then Name(id_of_string"_") else na in
let id = next_name_away na avoid in
((mkVar id)::subst,
id::avoid,
- add_named_decl (id,option_app (substl subst) c,
+ push_named (id,option_app (substl subst) c,
type_app (substl subst) t)
- sign))
- (rel_context env) ~init:([],ids_of_named_context sign0,sign0)
- in (subst, reset_with_named_context sign env)
+ env))
+ (rel_context env) ~init:([],ids_of_named_context (named_context env),env)
+ in (subst, (named_context_val env))
-let new_isevar isevars env src typ =
- let subst,env' = push_rel_context_to_named_context env in
+let new_evar evd env ?(src=(dummy_loc,InternalHole)) typ =
+ let subst,sign = push_rel_context_to_named_context env in
let typ' = substl subst typ in
let instance = make_evar_instance_with_rel env in
- let (sigma',evar) = new_isevar_sign env' isevars.evars typ' instance in
- isevars.evars <- sigma';
- isevars.history <- (fst (destEvar evar),src)::isevars.history;
- evar
+ new_evar_instance sign evd typ' ~src:src instance
+
+(* The same using side-effect *)
+let e_new_evar evd env ?(src=(dummy_loc,InternalHole)) ty =
+ let (evd',ev) = new_evar !evd env ~src:src ty in
+ evd := evd';
+ ev
+
+(*------------------------------------*
+ * operations on the evar constraints *
+ *------------------------------------*)
+
+(* Pb: defined Rels and Vars should not be considered as a pattern... *)
+let is_pattern inst =
+ let rec is_hopat l = function
+ [] -> true
+ | t :: tl ->
+ (isRel t or isVar t) && not (List.mem t l) && is_hopat (t::l) tl in
+ is_hopat [] (Array.to_list inst)
+
+let evar_well_typed_body evd ev evi body =
+ try
+ let env = evar_env evi in
+ let ty = evi.evar_concl in
+ Typing.check env (evars_of evd) body ty;
+ true
+ with e ->
+ pperrnl
+ (str "Ill-typed evar instantiation: " ++ fnl() ++
+ pr_evar_defs evd ++ fnl() ++
+ str "----> " ++ int ev ++ str " := " ++
+ print_constr body);
+ false
+
+let strict_inverse = false
+
+let inverse_instance env isevars ev evi inst rhs =
+ let subst = make_subst (evar_env evi) (Array.to_list inst) in
+ let subst = List.map (fun (x,y) -> (y,mkVar x)) subst in
+ let evd = ref isevars in
+ let error () =
+ error_not_clean env (evars_of !evd) ev rhs (evar_source ev !evd) in
+ let rec subs rigid k t =
+ match kind_of_term t with
+ | Rel i ->
+ if i<=k then t
+ else
+ (try List.assoc (mkRel (i-k)) subst
+ with Not_found ->
+ if rigid then error()
+ else if strict_inverse then
+ failwith "cannot solve pb yet"
+ else t)
+ | Var id ->
+ (try List.assoc t subst
+ with Not_found ->
+ if rigid then error()
+ else if
+ not strict_inverse &&
+ List.exists (fun (id',_,_) -> id=id') (evar_context evi)
+ then
+ failwith "cannot solve pb yet"
+ else t)
+ | Evar (ev,args) ->
+ if Evd.is_defined_evar !evd (ev,args) then
+ subs rigid k (existential_value (evars_of !evd) (ev,args))
+ else
+ let args' = Array.map (subs false k) args in
+ mkEvar (ev,args')
+ | _ -> map_constr_with_binders succ (subs rigid) k t in
+ let body = subs true 0 (nf_evar (evars_of isevars) rhs) in
+ (!evd,body)
+
+
+let is_defined_equation env evd (ev,inst) rhs =
+ is_pattern inst &&
+ not (occur_evar ev rhs) &&
+ try
+ let evi = Evd.map (evars_of evd) ev in
+ let (evd',body) = inverse_instance env evd ev evi inst rhs in
+ evar_well_typed_body evd' ev evi body
+ with Failure _ -> false
+
+
+(* Redefines an evar with a smaller context (i.e. it may depend on less
+ * variables) such that c becomes closed.
+ * Example: in [x:?1; y:(list ?2)] <?3>x=y /\ x=(nil bool)
+ * ?3 <-- ?1 no pb: env of ?3 is larger than ?1's
+ * ?1 <-- (list ?2) pb: ?2 may depend on x, but not ?1.
+ * What we do is that ?2 is defined by a new evar ?4 whose context will be
+ * a prefix of ?2's env, included in ?1's env. *)
+
+let do_restrict_hyps evd ev args =
+ let args = Array.to_list args in
+ let evi = Evd.map (evars_of !evd) ev in
+ let env = evar_env evi in
+ let hyps = evar_context evi in
+ let (sign,ncargs) = list_filter2 (fun _ a -> closed0 a) (hyps,args) in
+ (* No care is taken in case the evar type uses vars filtered out!
+ Is it important ? *)
+ let nc =
+ let env =
+ Sign.fold_named_context push_named sign ~init:(reset_context env) in
+ e_new_evar evd env ~src:(evar_source ev !evd) evi.evar_concl in
+ evd := Evd.evar_define ev nc !evd;
+ let (evn,_) = destEvar nc in
+ mkEvar(evn,Array.of_list ncargs)
+
+
+let need_restriction isevars args = not (array_for_all closed0 args)
+
+(* We try to instantiate the evar assuming the body won't depend
+ * on arguments that are not Rels or Vars, or appearing several times.
+ *)
+(* Note: error_not_clean should not be an error: it simply means that the
+ * conversion test that lead to the faulty call to [real_clean] should return
+ * false. The problem is that we won't get the right error message.
+ *)
+
+let real_clean env isevars ev evi args rhs =
+ let evd = ref isevars in
+ let subst = List.map (fun (x,y) -> (y,mkVar x)) (filter_unique args) in
+ let rec subs rigid k t =
+ match kind_of_term t with
+ | Rel i ->
+ if i<=k then t
+ else (try List.assoc (mkRel (i-k)) subst with Not_found -> t)
+ | Evar (ev,args) ->
+ if Evd.is_defined_evar !evd (ev,args) then
+ subs rigid k (existential_value (evars_of !evd) (ev,args))
+ else
+ let args' = Array.map (subs false k) args in
+ if need_restriction !evd args' then
+ do_restrict_hyps evd ev args'
+ else
+ mkEvar (ev,args')
+ | Var id ->
+ (try List.assoc t subst
+ with Not_found ->
+ if
+ not rigid
+ or List.exists (fun (id',_,_) -> id=id') (evar_context evi)
+ then t
+ else
+ error_not_clean env (evars_of !evd) ev rhs
+ (evar_source ev !evd))
+ | _ -> map_constr_with_binders succ (subs rigid) k t
+ in
+ let body = subs true 0 (nf_evar (evars_of isevars) rhs) in
+ if not (closed0 body)
+ then error_not_clean env (evars_of !evd) ev body (evar_source ev !evd);
+ (!evd,body)
(* [evar_define] solves the problem lhs = rhs when lhs is an uninstantiated
* evar, i.e. tries to find the body ?sp for lhs=mkEvar (sp,args)
@@ -331,30 +390,56 @@ let new_isevar isevars env src typ =
* ?1 would be instantiated by (le y y) but y is not in the scope of ?1
*)
-let evar_define env isevars (ev,argsv) rhs =
+(* env needed for error messages... *)
+let evar_define env (ev,argsv) rhs isevars =
if occur_evar ev rhs
then error_occur_check env (evars_of isevars) ev rhs;
let args = Array.to_list argsv in
- let evd = ise_map isevars ev in
+ let evi = Evd.map (evars_of isevars) ev in
(* the bindings to invert *)
- let worklist = make_subst (evar_env evd) args in
- let body = real_clean env isevars ev worklist rhs in
- ise_define isevars ev body;
- [ev]
+ let worklist = make_subst (evar_env evi) args in
+ let (isevars',body) = real_clean env isevars ev evi worklist rhs in
+ if occur_meta body then error "Meta cannot occur in evar body"
+ else
+ (* needed only if an inferred type *)
+ let body = refresh_universes body in
+(* Cannot strictly type instantiations since the unification algorithm
+ * does not unifies applications from left to right.
+ * e.g problem f x == g y yields x==y and f==g (in that order)
+ * Another problem is that type variables are evars of type Type
+ let _ =
+ try
+ let env = evar_env evi in
+ let ty = evi.evar_concl in
+ Typing.check env (evars_of isevars') body ty
+ with e ->
+ pperrnl
+ (str "Ill-typed evar instantiation: " ++ fnl() ++
+ pr_evar_defs isevars' ++ fnl() ++
+ str "----> " ++ int ev ++ str " := " ++
+ print_constr body);
+ raise e in*)
+ let isevars'' = Evd.evar_define ev body isevars' in
+ isevars'',[ev]
+
+
(*-------------------*)
(* Auxiliary functions for the conversion algorithms modulo evars
*)
-let has_undefined_isevars isevars t =
- try let _ = local_strong (whd_ise isevars.evars) t in false
+let has_undefined_evars isevars t =
+ try let _ = local_strong (whd_ise (evars_of isevars)) t in false
with Uninstantiated_evar _ -> true
+let is_ground_term isevars t =
+ not (has_undefined_evars isevars t)
+
let head_is_evar isevars =
let rec hrec k = match kind_of_term k with
- | Evar (n,_) -> not (Evd.is_defined isevars.evars n)
+ | Evar n -> not (Evd.is_defined_evar isevars n)
| App (f,_) -> hrec f
- | Cast (c,_) -> hrec c
+ | Cast (c,_,_) -> hrec c
| _ -> false
in
hrec
@@ -362,7 +447,7 @@ let head_is_evar isevars =
let rec is_eliminator c = match kind_of_term c with
| App _ -> true
| Case _ -> true
- | Cast (c,_) -> is_eliminator c
+ | Cast (c,_,_) -> is_eliminator c
| _ -> false
let head_is_embedded_evar isevars c =
@@ -373,7 +458,7 @@ let head_evar =
| Evar (ev,_) -> ev
| Case (_,_,c,_) -> hrec c
| App (c,_) -> hrec c
- | Cast (c,_) -> hrec c
+ | Cast (c,_,_) -> hrec c
| _ -> failwith "headconstant"
in
hrec
@@ -410,75 +495,70 @@ let status_changed lev (pbty,t1,t2) =
with Failure _ ->
try List.mem (head_evar t2) lev with Failure _ -> false
-let get_changed_pb isevars lev =
- let (pbs,pbs1) =
- List.fold_left
- (fun (pbs,pbs1) pb ->
- if status_changed lev pb then
- (pb::pbs,pbs1)
- else
- (pbs,pb::pbs1))
- ([],[])
- isevars.conv_pbs
- in
- isevars.conv_pbs <- pbs1;
- pbs
-
(* Solve pbs (?i x1..xn) = (?i y1..yn) which arises often in fixpoint
* definitions. We try to unify the xi with the yi pairwise. The pairs
* that don't unify are discarded (i.e. ?i is redefined so that it does not
* depend on these args). *)
let solve_refl conv_algo env isevars ev argsv1 argsv2 =
- if argsv1 = argsv2 then [] else
- let evd = Evd.map isevars.evars ev in
- let hyps = evd.evar_hyps in
- let (_,rsign) =
+ if argsv1 = argsv2 then (isevars,[]) else
+ let evd = Evd.map (evars_of isevars) ev in
+ let hyps = evar_context evd in
+ let (isevars',_,rsign) =
array_fold_left2
- (fun (sgn,rsgn) a1 a2 ->
- if conv_algo env isevars CONV a1 a2 then
- (List.tl sgn, add_named_decl (List.hd sgn) rsgn)
+ (fun (isevars,sgn,rsgn) a1 a2 ->
+ let (isevars',b) = conv_algo env isevars Reduction.CONV a1 a2 in
+ if b then
+ (isevars',List.tl sgn, add_named_decl (List.hd sgn) rsgn)
else
- (List.tl sgn, rsgn))
- (hyps,[]) argsv1 argsv2
+ (isevars,List.tl sgn, rsgn))
+ (isevars,hyps,[]) argsv1 argsv2
in
let nsign = List.rev rsign in
- let nargs = (Array.of_list (List.map mkVar (ids_of_named_context nsign))) in
- let newev = new_evar () in
- let info = { evar_concl = evd.evar_concl; evar_hyps = nsign;
- evar_body = Evar_empty } in
- isevars.evars <-
- Evd.define (Evd.add isevars.evars newev info) ev (mkEvar (newev,nargs));
- isevars.history <- (newev,evar_source ev isevars)::isevars.history;
- [ev]
+ let (evd',newev) =
+ let env =
+ Sign.fold_named_context push_named nsign ~init:(reset_context env) in
+ new_evar isevars env ~src:(evar_source ev isevars) evd.evar_concl in
+ let evd'' = Evd.evar_define ev newev evd' in
+ evd'', [ev]
(* Tries to solve problem t1 = t2.
- * Precondition: t1 is an uninstanciated evar
+ * Precondition: t1 is an uninstantiated evar
* Returns an optional list of evars that were instantiated, or None
* if the problem couldn't be solved. *)
(* Rq: uncomplete algorithm if pbty = CONV_X_LEQ ! *)
let solve_simple_eqn conv_algo env isevars (pbty,(n1,args1 as ev1),t2) =
- let t2 = nf_evar isevars.evars t2 in
- let lsp = match kind_of_term t2 with
- | Evar (n2,args2 as ev2)
- when not (Evd.is_defined isevars.evars n2) ->
- if n1 = n2 then
- solve_refl conv_algo env isevars n1 args1 args2
- else
- if Array.length args1 < Array.length args2 then
- evar_define env isevars ev2 (mkEvar ev1)
- else
- evar_define env isevars ev1 t2
- | _ ->
- evar_define env isevars ev1 t2 in
- let pbs = get_changed_pb isevars lsp in
- List.for_all (fun (pbty,t1,t2) -> conv_algo env isevars pbty t1 t2) pbs
+ try
+ let t2 = nf_evar (evars_of isevars) t2 in
+ let (isevars,lsp) = match kind_of_term t2 with
+ | Evar (n2,args2 as ev2) ->
+ if n1 = n2 then
+ solve_refl conv_algo env isevars n1 args1 args2
+ else
+ (try evar_define env ev1 t2 isevars
+ with e when precatchable_exception e ->
+ evar_define env ev2 (mkEvar ev1) isevars)
+(* if Array.length args1 < Array.length args2 then
+ evar_define env ev2 (mkEvar ev1) isevars
+ else
+ evar_define env ev1 t2 isevars*)
+ | _ ->
+ evar_define env ev1 t2 isevars in
+ let (isevars,pbs) = get_conv_pbs isevars (status_changed lsp) in
+ List.fold_left
+ (fun (isevars,b as p) (pbty,t1,t2) ->
+ if b then conv_algo env isevars pbty t1 t2 else p) (isevars,true)
+ pbs
+ with e when precatchable_exception e ->
+ (isevars,false)
(* Operations on value/type constraints *)
-type type_constraint = constr option
+type type_constraint_type = (int * int) option * constr
+type type_constraint = type_constraint_type option
+
type val_constraint = constr option
(* Old comment...
@@ -498,8 +578,14 @@ type val_constraint = constr option
(* The empty type constraint *)
let empty_tycon = None
+let mk_tycon_type c = (None, c)
+let mk_abstr_tycon_type n c = (Some (n, n), c) (* First component is initial abstraction, second
+ is current abstraction *)
+
(* Builds a type constraint *)
-let mk_tycon ty = Some ty
+let mk_tycon ty = Some (mk_tycon_type ty)
+
+let mk_abstr_tycon n ty = Some (mk_abstr_tycon_type n ty)
(* Constrains the value of a type *)
let empty_valcon = None
@@ -509,41 +595,98 @@ let mk_valcon c = Some c
(* Refining an evar to a product or a sort *)
-let refine_evar_as_arrow isevars ev =
- let (sigma,prod,evdom,evrng) = split_evar_to_arrow isevars.evars ev in
- evars_reset_evd sigma isevars;
- let hst = evar_source (fst ev) isevars in
- isevars.history <- (fst evrng,hst)::(fst evdom, hst)::isevars.history;
- (prod,evdom,evrng)
+(* Declaring any type to be in the sort Type shouldn't be harmful since
+ cumulativity now includes Prop and Set in Type...
+ It is, but that's not too bad *)
+let define_evar_as_abstraction abs evd (ev,args) =
+ let evi = Evd.map (evars_of evd) ev in
+ let evenv = evar_env evi in
+ let (evd1,dom) = new_evar evd evenv (new_Type()) in
+ let nvar =
+ next_ident_away (id_of_string "x")
+ (ids_of_named_context (evar_context evi)) in
+ let newenv = push_named (nvar, None, dom) evenv in
+ let (evd2,rng) =
+ new_evar evd1 newenv ~src:(evar_source ev evd1) (new_Type()) in
+ let prod = abs (Name nvar, dom, subst_var nvar rng) in
+ let evd3 = Evd.evar_define ev prod evd2 in
+ let evdom = fst (destEvar dom), args in
+ let evrng =
+ fst (destEvar rng), array_cons (mkRel 1) (Array.map (lift 1) args) in
+ let prod' = abs (Name nvar, mkEvar evdom, mkEvar evrng) in
+ (evd3,prod')
+
+let define_evar_as_arrow evd (ev,args) =
+ define_evar_as_abstraction (fun t -> mkProd t) evd (ev,args)
-let define_evar_as_arrow isevars ev =
- let (prod,_,_) = refine_evar_as_arrow isevars ev in
- prod
+let define_evar_as_lambda evd (ev,args) =
+ define_evar_as_abstraction (fun t -> mkLambda t) evd (ev,args)
let define_evar_as_sort isevars (ev,args) =
let s = new_Type () in
- let sigma' = Evd.define isevars.evars ev s in
- evars_reset_evd sigma' isevars;
- destSort s
+ Evd.evar_define ev s isevars, destSort s
+
+(* We don't try to guess in which sort the type should be defined, since
+ any type has type Type. May cause some trouble, but not so far... *)
+
+let judge_of_new_Type () = Typeops.judge_of_type (new_univ ())
(* Propagation of constraints through application and abstraction:
Given a type constraint on a functional term, returns the type
constraint on its domain and codomain. If the input constraint is
an evar instantiate it with the product of 2 new evars. *)
-let split_tycon loc env isevars = function
- | None -> Anonymous,None,None
- | Some c ->
- let sigma = evars_of isevars in
- let t = whd_betadeltaiota env sigma c in
+let split_tycon loc env isevars tycon =
+ let rec real_split c =
+ let sigma = evars_of isevars in
+ let t = whd_betadeltaiota env sigma c in
match kind_of_term t with
- | Prod (na,dom,rng) -> na, Some dom, Some rng
- | Evar (n,_ as ev) when not (Evd.is_defined isevars.evars n) ->
- let (_,evdom,evrng) = refine_evar_as_arrow isevars ev in
- Anonymous, Some (mkEvar evdom), Some (mkEvar evrng)
+ | Prod (na,dom,rng) -> isevars, (na, dom, rng)
+ | Evar ev when not (Evd.is_defined_evar isevars ev) ->
+ let (isevars',prod) = define_evar_as_arrow isevars ev in
+ let (_,dom,rng) = destProd prod in
+ isevars',(Anonymous, dom, rng)
| _ -> error_not_product_loc loc env sigma c
+ in
+ match tycon with
+ | None -> isevars,(Anonymous,None,None)
+ | Some (abs, c) ->
+ (match abs with
+ None ->
+ let isevars', (n, dom, rng) = real_split c in
+ isevars', (n, mk_tycon dom, mk_tycon rng)
+ | Some (init, cur) ->
+ if cur = 0 then
+ let isevars', (x, dom, rng) = real_split c in
+ isevars, (Anonymous,
+ Some (Some (init, 0), dom),
+ Some (Some (init, 0), rng))
+ else
+ isevars, (Anonymous, None, Some (Some (init, pred cur), c)))
+
+let valcon_of_tycon x =
+ match x with
+ | Some (None, t) -> Some t
+ | _ -> None
+
+let lift_abstr_tycon_type n (abs, t) =
+ match abs with
+ None -> raise (Invalid_argument "lift_abstr_tycon_type: not an abstraction")
+ | Some (init, abs) ->
+ let abs' = abs + n in
+ if abs' < 0 then raise (Invalid_argument "lift_abstr_tycon_type")
+ else (Some (init, abs'), t)
+
+let lift_tycon_type n (abs, t) = (abs, lift n t)
+let lift_tycon n = option_app (lift_tycon_type n)
+
+let pr_tycon_type env (abs, t) =
+ match abs with
+ None -> Termops.print_constr_env env t
+ | Some (init, cur) -> str "Abstract (" ++ int init ++ str "," ++ int cur ++ str ") " ++ Termops.print_constr_env env t
+
+let pr_tycon env = function
+ None -> str "None"
+ | Some t -> pr_tycon_type env t
-let valcon_of_tycon x = x
-
-let lift_tycon = option_app (lift 1)
diff --git a/pretyping/evarutil.mli b/pretyping/evarutil.mli
index 011d2a92..7429cd16 100644
--- a/pretyping/evarutil.mli
+++ b/pretyping/evarutil.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evarutil.mli,v 1.33.2.2 2004/07/16 19:30:44 herbelin Exp $ i*)
+(*i $Id: evarutil.mli 8695 2006-04-10 16:33:52Z msozeau $ i*)
(*i*)
open Util
@@ -21,77 +21,132 @@ open Reductionops
(*s This modules provides useful functions for unification modulo evars *)
-(* [whd_ise] raise [Uninstantiated_evar] if an evar remains uninstantiated; *)
-(* *[whd_evar]* and *[nf_evar]* leave uninstantiated evar as is *)
+(***********************************************************)
+(* Metas *)
-val nf_evar : evar_map -> constr -> constr
-val j_nf_evar : evar_map -> unsafe_judgment -> unsafe_judgment
-val jl_nf_evar :
- evar_map -> unsafe_judgment list -> unsafe_judgment list
-val jv_nf_evar :
- evar_map -> unsafe_judgment array -> unsafe_judgment array
-val tj_nf_evar :
- evar_map -> unsafe_type_judgment -> unsafe_type_judgment
+(* [new_meta] is a generator of unique meta variables *)
+val new_meta : unit -> metavariable
+val mk_new_meta : unit -> constr
-(* Replacing all evars *)
-exception Uninstantiated_evar of existential_key
-val whd_ise : evar_map -> constr -> constr
-val whd_castappevar : evar_map -> constr -> constr
+(* [new_untyped_evar] is a generator of unique evar keys *)
+val new_untyped_evar : unit -> existential_key
+
+(***********************************************************)
+(* Creating a fresh evar given their type and context *)
+val new_evar :
+ evar_defs -> env -> ?src:loc * hole_kind -> types -> evar_defs * constr
+(* the same with side-effects *)
+val e_new_evar :
+ evar_defs ref -> env -> ?src:loc * hole_kind -> types -> constr
+
+(* Create a fresh evar in a context different from its definition context:
+ [new_evar_instance sign evd ty inst] creates a new evar of context
+ [sign] and type [ty], [inst] is a mapping of the evar context to
+ the context where the evar should occur. This means that the terms
+ of [inst] are typed in the occurrence context and their type (seen
+ as a telescope) is [sign] *)
+val new_evar_instance :
+ named_context_val -> evar_defs -> types -> ?src:loc * hole_kind ->
+ constr list -> evar_defs * constr
-(* Creating new existential variables *)
-val new_evar : unit -> evar
-val new_evar_in_sign : env -> constr
+(***********************************************************)
+(* Instanciate evars *)
-val evar_env : evar_info -> env
+(* suspicious env ? *)
+val evar_define :
+ env -> existential -> constr -> evar_defs -> evar_defs * evar list
-type evar_defs
-val evars_of : evar_defs -> evar_map
-val create_evar_defs : evar_map -> evar_defs
-val evars_reset_evd : evar_map -> evar_defs -> unit
-val evar_source : existential_key -> evar_defs -> loc * hole_kind
-type evar_constraint = conv_pb * constr * constr
-val add_conv_pb : evar_defs -> evar_constraint -> unit
+(***********************************************************)
+(* Evars/Metas switching... *)
-val is_defined_evar : evar_defs -> existential -> bool
-val ise_try : evar_defs -> (unit -> bool) list -> bool
-val ise_undefined : evar_defs -> constr -> bool
-val has_undefined_isevars : evar_defs -> constr -> bool
+(* [evars_to_metas] generates new metavariables for each non dependent
+ existential and performs the replacement in the given constr; it also
+ returns the evar_map extended with dependent evars *)
+val evars_to_metas : evar_map -> open_constr -> (evar_map * constr)
-val new_isevar_sign :
- Environ.env -> Evd.evar_map -> Term.constr -> Term.constr list ->
- Evd.evar_map * Term.constr
+val non_instantiated : evar_map -> (evar * evar_info) list
-val new_isevar : evar_defs -> env -> loc * hole_kind -> constr -> constr
+(***********************************************************)
+(* Unification utils *)
+val is_ground_term : evar_defs -> constr -> bool
val is_eliminator : constr -> bool
val head_is_embedded_evar : evar_defs -> constr -> bool
val solve_simple_eqn :
- (env -> evar_defs -> conv_pb -> constr -> constr -> bool)
- -> env -> evar_defs -> conv_pb * existential * constr -> bool
+ (env -> evar_defs -> conv_pb -> constr -> constr -> evar_defs * bool)
+ -> env -> evar_defs -> conv_pb * existential * constr ->
+ evar_defs * bool
-val define_evar_as_arrow : evar_defs -> existential -> types
-val define_evar_as_sort : evar_defs -> existential -> sorts
+val define_evar_as_arrow : evar_defs -> existential -> evar_defs * types
+val define_evar_as_lambda : evar_defs -> existential -> evar_defs * types
+val define_evar_as_sort : evar_defs -> existential -> evar_defs * sorts
+(***********************************************************)
(* Value/Type constraints *)
-val new_Type_sort : unit -> sorts
-val new_Type : unit -> constr
val judge_of_new_Type : unit -> unsafe_judgment
-val refresh_universes : types -> types
-type type_constraint = constr option
+type type_constraint_type = (int * int) option * constr
+type type_constraint = type_constraint_type option
+
type val_constraint = constr option
val empty_tycon : type_constraint
+val mk_tycon_type : constr -> type_constraint_type
+val mk_abstr_tycon_type : int -> constr -> type_constraint_type
val mk_tycon : constr -> type_constraint
+val mk_abstr_tycon : int -> constr -> type_constraint
val empty_valcon : val_constraint
val mk_valcon : constr -> val_constraint
val split_tycon :
loc -> env -> evar_defs -> type_constraint ->
- name * type_constraint * type_constraint
+ evar_defs * (name * type_constraint * type_constraint)
val valcon_of_tycon : type_constraint -> val_constraint
-val lift_tycon : type_constraint -> type_constraint
+val lift_abstr_tycon_type : int -> type_constraint_type -> type_constraint_type
+
+val lift_tycon_type : int -> type_constraint_type -> type_constraint_type
+val lift_tycon : int -> type_constraint -> type_constraint
+
+(***********************************************************)
+
+(* [whd_ise] raise [Uninstantiated_evar] if an evar remains uninstantiated; *)
+(* *[whd_evar]* and *[nf_evar]* leave uninstantiated evar as is *)
+
+val nf_evar : evar_map -> constr -> constr
+val j_nf_evar : evar_map -> unsafe_judgment -> unsafe_judgment
+val jl_nf_evar :
+ evar_map -> unsafe_judgment list -> unsafe_judgment list
+val jv_nf_evar :
+ evar_map -> unsafe_judgment array -> unsafe_judgment array
+val tj_nf_evar :
+ evar_map -> unsafe_type_judgment -> unsafe_type_judgment
+
+val nf_evar_info : evar_map -> evar_info -> evar_info
+val nf_evars : evar_map -> evar_map
+
+(* Same for evar defs *)
+val nf_isevar : evar_defs -> constr -> constr
+val j_nf_isevar : evar_defs -> unsafe_judgment -> unsafe_judgment
+val jl_nf_isevar :
+ evar_defs -> unsafe_judgment list -> unsafe_judgment list
+val jv_nf_isevar :
+ evar_defs -> unsafe_judgment array -> unsafe_judgment array
+val tj_nf_isevar :
+ evar_defs -> unsafe_type_judgment -> unsafe_type_judgment
+
+val nf_evar_defs : evar_defs -> evar_defs
+
+(* Replacing all evars *)
+exception Uninstantiated_evar of existential_key
+val whd_ise : evar_map -> constr -> constr
+val whd_castappevar : evar_map -> constr -> constr
+
+(*********************************************************************)
+(* debug pretty-printer: *)
+
+val pr_tycon_type : env -> type_constraint_type -> Pp.std_ppcmds
+val pr_tycon : env -> type_constraint -> Pp.std_ppcmds
diff --git a/pretyping/evd.ml b/pretyping/evd.ml
index 7a3e7c02..c9f771c9 100644
--- a/pretyping/evd.ml
+++ b/pretyping/evd.ml
@@ -6,12 +6,18 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: evd.ml,v 1.3.2.1 2004/07/16 19:30:44 herbelin Exp $ *)
+(* $Id: evd.ml 8688 2006-04-07 15:08:12Z msozeau $ *)
+open Pp
open Util
open Names
+open Nameops
open Term
+open Termops
open Sign
+open Environ
+open Libnames
+open Mod_subst
(* The type of mappings for existential variables *)
@@ -23,12 +29,20 @@ type evar_body =
type evar_info = {
evar_concl : constr;
- evar_hyps : named_context;
+ evar_hyps : named_context_val;
evar_body : evar_body}
+let evar_context evi = named_context_of_val evi.evar_hyps
+
+let eq_evar_info ei1 ei2 =
+ ei1 == ei2 ||
+ eq_constr ei1.evar_concl ei2.evar_concl &&
+ eq_named_context_val (ei1.evar_hyps) (ei2.evar_hyps) &&
+ ei1.evar_body = ei2.evar_body
+
module Evarmap = Intmap
-type evar_map = evar_info Evarmap.t
+type evar_map1 = evar_info Evarmap.t
let empty = Evarmap.empty
@@ -38,28 +52,21 @@ let map evc k = Evarmap.find k evc
let rmv evc k = Evarmap.remove k evc
let remap evc k i = Evarmap.add k i evc
let in_dom evc k = Evarmap.mem k evc
+let fold = Evarmap.fold
let add evd ev newinfo = Evarmap.add ev newinfo evd
let define evd ev body =
- let oldinfo = map evd ev in
+ let oldinfo =
+ try map evd ev
+ with Not_found -> error "Evd.define: cannot define undeclared evar" in
let newinfo =
{ evar_concl = oldinfo.evar_concl;
evar_hyps = oldinfo.evar_hyps;
- evar_body = Evar_defined body}
- in
+ evar_body = Evar_defined body} in
match oldinfo.evar_body with
| Evar_empty -> Evarmap.add ev newinfo evd
- | _ -> anomaly "cannot define an isevar twice"
-
-(* The list of non-instantiated existential declarations *)
-
-let non_instantiated sigma =
- let listev = to_list sigma in
- List.fold_left
- (fun l ((ev,evd) as d) ->
- if evd.evar_body = Evar_empty then (d::l) else l)
- [] listev
+ | _ -> anomaly "Evd.define: cannot define an isevar twice"
let is_evar sigma ev = in_dom sigma ev
@@ -68,7 +75,490 @@ let is_defined sigma ev =
not (info.evar_body = Evar_empty)
let evar_body ev = ev.evar_body
+let evar_env evd = Global.env_of_context evd.evar_hyps
let string_of_existential ev = "?" ^ string_of_int ev
let existential_of_int ev = ev
+
+(*******************************************************************)
+(* Formerly Instantiate module *)
+
+let is_id_inst inst =
+ let is_id (id,c) = match kind_of_term c with
+ | Var id' -> id = id'
+ | _ -> false
+ in
+ List.for_all is_id inst
+
+(* Vérifier que les instances des let-in sont compatibles ?? *)
+let instantiate_sign_including_let sign args =
+ let rec instrec = function
+ | ((id,b,_) :: sign, c::args) -> (id,c) :: (instrec (sign,args))
+ | ([],[]) -> []
+ | ([],_) | (_,[]) ->
+ anomaly "Signature and its instance do not match"
+ in
+ instrec (sign,args)
+
+let instantiate_evar sign c args =
+ let inst = instantiate_sign_including_let sign args in
+ if is_id_inst inst then
+ c
+ else
+ replace_vars inst c
+
+(* Existentials. *)
+
+let existential_type sigma (n,args) =
+ let info =
+ try map sigma n
+ with Not_found ->
+ anomaly ("Evar "^(string_of_existential n)^" was not declared") in
+ let hyps = evar_context info in
+ instantiate_evar hyps info.evar_concl (Array.to_list args)
+
+exception NotInstantiatedEvar
+
+let existential_value sigma (n,args) =
+ let info = map sigma n in
+ let hyps = evar_context info in
+ match evar_body info with
+ | Evar_defined c ->
+ instantiate_evar hyps c (Array.to_list args)
+ | Evar_empty ->
+ raise NotInstantiatedEvar
+
+let existential_opt_value sigma ev =
+ try Some (existential_value sigma ev)
+ with NotInstantiatedEvar -> None
+
+(*******************************************************************)
+(* Constraints for sort variables *)
+(*******************************************************************)
+
+type sort_var = Univ.universe
+
+type sort_constraint =
+ | DefinedSort of sorts (* instantiated sort var *)
+ | SortVar of sort_var list * sort_var list (* (leq,geq) *)
+ | EqSort of sort_var
+
+module UniverseOrdered = struct
+ type t = Univ.universe
+ let compare = Pervasives.compare
+end
+module UniverseMap = Map.Make(UniverseOrdered)
+
+type sort_constraints = sort_constraint UniverseMap.t
+
+let rec canonical_find u scstr =
+ match UniverseMap.find u scstr with
+ EqSort u' -> canonical_find u' scstr
+ | c -> (u,c)
+
+let whd_sort_var scstr t =
+ match kind_of_term t with
+ Sort(Type u) ->
+ (try
+ match canonical_find u scstr with
+ _, DefinedSort s -> mkSort s
+ | _ -> t
+ with Not_found -> t)
+ | _ -> t
+
+let rec set_impredicative u s scstr =
+ match UniverseMap.find u scstr with
+ | DefinedSort s' ->
+ if family_of_sort s = family_of_sort s' then scstr
+ else failwith "sort constraint inconsistency"
+ | EqSort u' ->
+ UniverseMap.add u (DefinedSort s) (set_impredicative u' s scstr)
+ | SortVar(_,ul) ->
+ (* also set sorts lower than u as impredicative *)
+ UniverseMap.add u (DefinedSort s)
+ (List.fold_left (fun g u' -> set_impredicative u' s g) scstr ul)
+
+let rec set_predicative u s scstr =
+ match UniverseMap.find u scstr with
+ | DefinedSort s' ->
+ if family_of_sort s = family_of_sort s' then scstr
+ else failwith "sort constraint inconsistency"
+ | EqSort u' ->
+ UniverseMap.add u (DefinedSort s) (set_predicative u' s scstr)
+ | SortVar(ul,_) ->
+ UniverseMap.add u (DefinedSort s)
+ (List.fold_left (fun g u' -> set_impredicative u' s g) scstr ul)
+
+let var_of_sort = function
+ Type u -> u
+ | _ -> assert false
+
+let is_sort_var s scstr =
+ match s with
+ Type u ->
+ (try
+ match canonical_find u scstr with
+ _, DefinedSort _ -> false
+ | _ -> true
+ with Not_found -> false)
+ | _ -> false
+
+let new_sort_var cstr =
+ let u = Termops.new_univ() in
+ (u, UniverseMap.add u (SortVar([],[])) cstr)
+
+
+let set_leq_sort (u1,(leq1,geq1)) (u2,(leq2,geq2)) scstr =
+ let rec search_rec (is_b, betw, not_betw) u1 =
+ if List.mem u1 betw then (true, betw, not_betw)
+ else if List.mem u1 not_betw then (is_b, betw, not_betw)
+ else if u1 = u2 then (true, u1::betw,not_betw) else
+ match UniverseMap.find u1 scstr with
+ EqSort u1' -> search_rec (is_b,betw,not_betw) u1'
+ | SortVar(leq,_) ->
+ let (is_b',betw',not_betw') =
+ List.fold_left search_rec (false,betw,not_betw) leq in
+ if is_b' then (true, u1::betw', not_betw')
+ else (false, betw', not_betw')
+ | DefinedSort _ -> (false,betw,u1::not_betw) in
+ let (is_betw,betw,_) = search_rec (false, [], []) u1 in
+ if is_betw then
+ UniverseMap.add u1 (SortVar(leq1@leq2,geq1@geq2))
+ (List.fold_left
+ (fun g u -> UniverseMap.add u (EqSort u1) g) scstr betw)
+ else
+ UniverseMap.add u1 (SortVar(u2::leq1,geq1))
+ (UniverseMap.add u2 (SortVar(leq2, u1::geq2)) scstr)
+
+let set_leq s1 s2 scstr =
+ let u1 = var_of_sort s1 in
+ let u2 = var_of_sort s2 in
+ let (cu1,c1) = canonical_find u1 scstr in
+ let (cu2,c2) = canonical_find u2 scstr in
+ if cu1=cu2 then scstr
+ else
+ match c1,c2 with
+ (EqSort _, _ | _, EqSort _) -> assert false
+ | SortVar(leq1,geq1), SortVar(leq2,geq2) ->
+ set_leq_sort (cu1,(leq1,geq1)) (cu2,(leq2,geq2)) scstr
+ | _, DefinedSort(Prop _ as s) -> set_impredicative u1 s scstr
+ | _, DefinedSort(Type _) -> scstr
+ | DefinedSort(Type _ as s), _ -> set_predicative u2 s scstr
+ | DefinedSort(Prop _), _ -> scstr
+
+let set_sort_variable s1 s2 scstr =
+ let u = var_of_sort s1 in
+ match s2 with
+ Prop _ -> set_impredicative u s2 scstr
+ | Type _ -> set_predicative u s2 scstr
+
+let pr_sort_cstrs g =
+ let l = UniverseMap.fold (fun u c l -> (u,c)::l) g [] in
+ str "SORT CONSTRAINTS:" ++ fnl() ++
+ prlist_with_sep fnl (fun (u,c) ->
+ match c with
+ EqSort u' -> Univ.pr_uni u ++ str" == " ++ Univ.pr_uni u'
+ | DefinedSort s -> Univ.pr_uni u ++ str " := " ++ print_sort s
+ | SortVar(leq,geq) ->
+ str"[" ++ hov 0 (prlist_with_sep spc Univ.pr_uni geq) ++
+ str"] <= "++ Univ.pr_uni u ++ brk(0,0) ++ str"<= [" ++
+ hov 0 (prlist_with_sep spc Univ.pr_uni leq) ++ str"]")
+ l
+
+type evar_map = evar_map1 * sort_constraints
+let empty = empty, UniverseMap.empty
+let add (sigma,sm) k v = (add sigma k v, sm)
+let dom (sigma,_) = dom sigma
+let map (sigma,_) = map sigma
+let rmv (sigma,sm) k = (rmv sigma k, sm)
+let remap (sigma,sm) k v = (remap sigma k v, sm)
+let in_dom (sigma,_) = in_dom sigma
+let to_list (sigma,_) = to_list sigma
+let fold f (sigma,_) = fold f sigma
+let define (sigma,sm) k v = (define sigma k v, sm)
+let is_evar (sigma,_) = is_evar sigma
+let is_defined (sigma,_) = is_defined sigma
+let existential_value (sigma,_) = existential_value sigma
+let existential_type (sigma,_) = existential_type sigma
+let existential_opt_value (sigma,_) = existential_opt_value sigma
+
+(*******************************************************************)
+type open_constr = evar_map * constr
+
+(*******************************************************************)
+(* The type constructor ['a sigma] adds an evar map to an object of
+ type ['a] *)
+type 'a sigma = {
+ it : 'a ;
+ sigma : evar_map}
+
+let sig_it x = x.it
+let sig_sig x = x.sigma
+
+(*******************************************************************)
+(* Metamaps *)
+
+(*******************************************************************)
+(* Constraints for existential variables *)
+(*******************************************************************)
+
+type 'a freelisted = {
+ rebus : 'a;
+ freemetas : Intset.t }
+
+(* Collects all metavars appearing in a constr *)
+let metavars_of c =
+ let rec collrec acc c =
+ match kind_of_term c with
+ | Meta mv -> Intset.add mv acc
+ | _ -> fold_constr collrec acc c
+ in
+ collrec Intset.empty c
+
+let mk_freelisted c =
+ { rebus = c; freemetas = metavars_of c }
+
+let map_fl f cfl = { cfl with rebus=f cfl.rebus }
+
+
+(* Clausal environments *)
+
+type clbinding =
+ | Cltyp of name * constr freelisted
+ | Clval of name * constr freelisted * constr freelisted
+
+let map_clb f = function
+ | Cltyp (na,cfl) -> Cltyp (na,map_fl f cfl)
+ | Clval (na,cfl1,cfl2) -> Clval (na,map_fl f cfl1,map_fl f cfl2)
+
+(* name of defined is erased (but it is pretty-printed) *)
+let clb_name = function
+ Cltyp(na,_) -> (na,false)
+ | Clval (na,_,_) -> (na,true)
+
+(***********************)
+
+module Metaset = Intset
+
+let meta_exists p s = Metaset.fold (fun x b -> b || (p x)) s false
+
+module Metamap = Intmap
+
+let metamap_to_list m =
+ Metamap.fold (fun n v l -> (n,v)::l) m []
+
+(*************************)
+(* Unification state *)
+
+type hole_kind =
+ | ImplicitArg of global_reference * (int * identifier option)
+ | BinderType of name
+ | QuestionMark
+ | CasesType
+ | InternalHole
+ | TomatchTypeParameter of inductive * int
+
+type conv_pb = Reduction.conv_pb
+type evar_constraint = conv_pb * constr * constr
+type evar_defs =
+ { evars : evar_map;
+ conv_pbs : evar_constraint list;
+ history : (existential_key * (loc * hole_kind)) list;
+ metas : clbinding Metamap.t }
+
+let subst_evar_defs sub evd =
+ { evd with
+ conv_pbs =
+ List.map (fun (k,t1,t2) ->(k,subst_mps sub t1,subst_mps sub t2))
+ evd.conv_pbs;
+ metas = Metamap.map (map_clb (subst_mps sub)) evd.metas }
+
+let create_evar_defs sigma =
+ { evars=sigma; conv_pbs=[]; history=[]; metas=Metamap.empty }
+let evars_of d = d.evars
+let evars_reset_evd evd d = {d with evars = evd}
+let reset_evd (sigma,mmap) d = {d with evars = sigma; metas=mmap}
+let add_conv_pb pb d =
+(* let (pbty,c1,c2) = pb in
+ pperrnl
+ (Termops.print_constr c1 ++
+ (if pbty=Reduction.CUMUL then str " <="++ spc()
+ else str" =="++spc()) ++
+ Termops.print_constr c2);*)
+ {d with conv_pbs = pb::d.conv_pbs}
+let evar_source ev d =
+ try List.assoc ev d.history
+ with Not_found -> (dummy_loc, InternalHole)
+
+(* define the existential of section path sp as the constr body *)
+let evar_define sp body isevars =
+ {isevars with evars = define isevars.evars sp body}
+
+let evar_declare hyps evn ty ?(src=(dummy_loc,InternalHole)) evd =
+ { evd with
+ evars = add evd.evars evn
+ {evar_hyps=hyps; evar_concl=ty; evar_body=Evar_empty};
+ history = (evn,src)::evd.history }
+
+let is_defined_evar isevars (n,_) = is_defined isevars.evars n
+
+(* Does k corresponds to an (un)defined existential ? *)
+let is_undefined_evar isevars c = match kind_of_term c with
+ | Evar ev -> not (is_defined_evar isevars ev)
+ | _ -> false
+
+let undefined_evars isevars =
+ let evd =
+ fold (fun ev evi sigma -> if evi.evar_body = Evar_empty then
+ add sigma ev evi else sigma)
+ isevars.evars empty
+ in
+ { isevars with evars = evd }
+
+(* extracts conversion problems that satisfy predicate p *)
+(* Note: conv_pbs not satisying p are stored back in reverse order *)
+let get_conv_pbs isevars p =
+ let (pbs,pbs1) =
+ List.fold_left
+ (fun (pbs,pbs1) pb ->
+ if p pb then
+ (pb::pbs,pbs1)
+ else
+ (pbs,pb::pbs1))
+ ([],[])
+ isevars.conv_pbs
+ in
+ {isevars with conv_pbs = pbs1},
+ pbs
+
+
+(**********************************************************)
+(* Sort variables *)
+
+let new_sort_variable (sigma,sm) =
+ let (u,scstr) = new_sort_var sm in
+ (Type u,(sigma,scstr))
+let is_sort_variable (_,sm) s =
+ is_sort_var s sm
+let whd_sort_variable (_,sm) t = whd_sort_var sm t
+let set_leq_sort_variable (sigma,sm) u1 u2 =
+ (sigma, set_leq u1 u2 sm)
+let define_sort_variable (sigma,sm) u s =
+ (sigma, set_sort_variable u s sm)
+let pr_sort_constraints (_,sm) = pr_sort_cstrs sm
+
+(**********************************************************)
+(* Accessing metas *)
+
+let meta_list evd = metamap_to_list evd.metas
+
+let meta_defined evd mv =
+ match Metamap.find mv evd.metas with
+ | Clval _ -> true
+ | Cltyp _ -> false
+
+let meta_fvalue evd mv =
+ match Metamap.find mv evd.metas with
+ | Clval(_,b,_) -> b
+ | Cltyp _ -> anomaly "meta_fvalue: meta has no value"
+
+let meta_ftype evd mv =
+ match Metamap.find mv evd.metas with
+ | Cltyp (_,b) -> b
+ | Clval(_,_,b) -> b
+
+let meta_declare mv v ?(name=Anonymous) evd =
+ { evd with metas = Metamap.add mv (Cltyp(name,mk_freelisted v)) evd.metas }
+
+let meta_assign mv v evd =
+ match Metamap.find mv evd.metas with
+ Cltyp(na,ty) ->
+ { evd with
+ metas = Metamap.add mv (Clval(na,mk_freelisted v, ty)) evd.metas }
+ | _ -> anomaly "meta_assign: already defined"
+
+(* If the meta is defined then forget its name *)
+let meta_name evd mv =
+ try
+ let (na,def) = clb_name (Metamap.find mv evd.metas) in
+ if def then Anonymous else na
+ with Not_found -> Anonymous
+
+let meta_with_name evd id =
+ let na = Name id in
+ let (mvl,mvnodef) =
+ Metamap.fold
+ (fun n clb (l1,l2 as l) ->
+ let (na',def) = clb_name clb in
+ if na = na' then if def then (n::l1,l2) else (n::l1,n::l2)
+ else l)
+ evd.metas ([],[]) in
+ match mvnodef, mvl with
+ | _,[] ->
+ errorlabstrm "Evd.meta_with_name"
+ (str"No such bound variable " ++ pr_id id)
+ | ([n],_|_,[n]) ->
+ n
+ | _ ->
+ errorlabstrm "Evd.meta_with_name"
+ (str "Binder name \"" ++ pr_id id ++
+ str"\" occurs more than once in clause")
+
+
+let meta_merge evd1 evd2 =
+ {evd2 with
+ metas = List.fold_left (fun m (n,v) -> Metamap.add n v m)
+ evd2.metas (metamap_to_list evd1.metas) }
+
+
+(**********************************************************)
+(* Pretty-printing *)
+
+let pr_meta_map mmap =
+ let pr_name = function
+ Name id -> str"[" ++ pr_id id ++ str"]"
+ | _ -> mt() in
+ let pr_meta_binding = function
+ | (mv,Cltyp (na,b)) ->
+ hov 0
+ (pr_meta mv ++ pr_name na ++ str " : " ++
+ print_constr b.rebus ++ fnl ())
+ | (mv,Clval(na,b,_)) ->
+ hov 0
+ (pr_meta mv ++ pr_name na ++ str " := " ++
+ print_constr b.rebus ++ fnl ())
+ in
+ prlist pr_meta_binding (metamap_to_list mmap)
+
+let pr_idl idl = prlist_with_sep pr_spc pr_id idl
+
+let pr_evar_info evi =
+ let phyps = pr_idl (List.rev (ids_of_named_context (evar_context evi))) in
+ let pty = print_constr evi.evar_concl in
+ let pb =
+ match evi.evar_body with
+ | Evar_empty -> mt ()
+ | Evar_defined c -> spc() ++ str"=> " ++ print_constr c
+ in
+ hov 2 (str"[" ++ phyps ++ spc () ++ str"|- " ++ pty ++ pb ++ str"]")
+
+let pr_evar_map sigma =
+ h 0
+ (prlist_with_sep pr_fnl
+ (fun (ev,evi) ->
+ h 0 (str(string_of_existential ev)++str"=="++ pr_evar_info evi))
+ (to_list sigma))
+
+let pr_evar_defs evd =
+ let pp_evm =
+ if evd.evars = empty then mt() else
+ str"EVARS:"++brk(0,1)++pr_evar_map evd.evars++fnl() in
+ let n = List.length evd.conv_pbs in
+ let cstrs =
+ if n=0 then mt() else
+ str"=> " ++ int n ++ str" constraints" ++ fnl() ++ fnl() in
+ let pp_met =
+ if evd.metas = Metamap.empty then mt() else
+ str"METAS:"++brk(0,1)++pr_meta_map evd.metas in
+ v 0 (pp_evm ++ cstrs ++ pp_met)
diff --git a/pretyping/evd.mli b/pretyping/evd.mli
index f66667aa..40ecce6e 100644
--- a/pretyping/evd.mli
+++ b/pretyping/evd.mli
@@ -6,12 +6,15 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: evd.mli,v 1.3.2.1 2004/07/16 19:30:44 herbelin Exp $ i*)
+(*i $Id: evd.mli 8688 2006-04-07 15:08:12Z msozeau $ i*)
(*i*)
+open Util
open Names
open Term
open Sign
+open Libnames
+open Mod_subst
(*i*)
(* The type of mappings for existential variables.
@@ -28,9 +31,11 @@ type evar_body =
type evar_info = {
evar_concl : constr;
- evar_hyps : named_context;
+ evar_hyps : Environ.named_context_val;
evar_body : evar_body}
+val eq_evar_info : evar_info -> evar_info -> bool
+val evar_context : evar_info -> named_context
type evar_map
val empty : evar_map
@@ -43,15 +48,127 @@ val rmv : evar_map -> evar -> evar_map
val remap : evar_map -> evar -> evar_info -> evar_map
val in_dom : evar_map -> evar -> bool
val to_list : evar_map -> (evar * evar_info) list
+val fold : (evar -> evar_info -> 'a -> 'a) -> evar_map -> 'a -> 'a
val define : evar_map -> evar -> constr -> evar_map
-val non_instantiated : evar_map -> (evar * evar_info) list
val is_evar : evar_map -> evar -> bool
val is_defined : evar_map -> evar -> bool
val evar_body : evar_info -> evar_body
+val evar_env : evar_info -> Environ.env
val string_of_existential : evar -> string
val existential_of_int : int -> evar
+
+(*s [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
+ no body and [Not_found] if it does not exist in [sigma] *)
+
+exception NotInstantiatedEvar
+val existential_value : evar_map -> existential -> constr
+val existential_type : evar_map -> existential -> types
+val existential_opt_value : evar_map -> existential -> constr option
+
+(*********************************************************************)
+(* constr with holes *)
+type open_constr = evar_map * constr
+
+(*********************************************************************)
+(* The type constructor ['a sigma] adds an evar map to an object of
+ type ['a] *)
+type 'a sigma = {
+ it : 'a ;
+ sigma : evar_map}
+
+val sig_it : 'a sigma -> 'a
+val sig_sig : 'a sigma -> evar_map
+
+(*********************************************************************)
+(* Meta map *)
+
+module Metaset : Set.S with type elt = metavariable
+
+val meta_exists : (metavariable -> bool) -> Metaset.t -> bool
+
+type 'a freelisted = {
+ rebus : 'a;
+ freemetas : Metaset.t }
+
+val mk_freelisted : constr -> constr freelisted
+val map_fl : ('a -> 'b) -> 'a freelisted -> 'b freelisted
+
+type clbinding =
+ | Cltyp of name * constr freelisted
+ | Clval of name * constr freelisted * constr freelisted
+
+val map_clb : (constr -> constr) -> clbinding -> clbinding
+
+(*********************************************************************)
+(* Unification state *)
+type evar_defs
+
+(* Substitution is not applied to the [evar_map] *)
+val subst_evar_defs : substitution -> evar_defs -> evar_defs
+
+(* create an [evar_defs] with empty meta map: *)
+val create_evar_defs : evar_map -> evar_defs
+val evars_of : evar_defs -> evar_map
+val evars_reset_evd : evar_map -> evar_defs -> evar_defs
+
+(* Evars *)
+type hole_kind =
+ | ImplicitArg of global_reference * (int * identifier option)
+ | BinderType of name
+ | QuestionMark
+ | CasesType
+ | InternalHole
+ | TomatchTypeParameter of inductive * int
+val is_defined_evar : evar_defs -> existential -> bool
+val is_undefined_evar : evar_defs -> constr -> bool
+val undefined_evars : evar_defs -> evar_defs
+val evar_declare :
+ Environ.named_context_val -> evar -> types -> ?src:loc * hole_kind ->
+ evar_defs -> evar_defs
+val evar_define : evar -> constr -> evar_defs -> evar_defs
+val evar_source : existential_key -> evar_defs -> loc * hole_kind
+
+(* Unification constraints *)
+type conv_pb = Reduction.conv_pb
+type evar_constraint = conv_pb * constr * constr
+val add_conv_pb : evar_constraint -> evar_defs -> evar_defs
+val get_conv_pbs : evar_defs -> (evar_constraint -> bool) ->
+ evar_defs * evar_constraint list
+
+(* Metas *)
+val meta_list : evar_defs -> (metavariable * clbinding) list
+val meta_defined : evar_defs -> metavariable -> bool
+(* [meta_fvalue] raises [Not_found] if meta not in map or [Anomaly] if
+ meta has no value *)
+val meta_fvalue : evar_defs -> metavariable -> constr freelisted
+val meta_ftype : evar_defs -> metavariable -> constr freelisted
+val meta_name : evar_defs -> metavariable -> name
+val meta_with_name : evar_defs -> identifier -> metavariable
+val meta_declare :
+ metavariable -> types -> ?name:name -> evar_defs -> evar_defs
+val meta_assign : metavariable -> constr -> evar_defs -> evar_defs
+
+(* [meta_merge evd1 evd2] returns [evd2] extended with the metas of [evd1] *)
+val meta_merge : evar_defs -> evar_defs -> evar_defs
+
+(**********************************************************)
+(* Sort variables *)
+
+val new_sort_variable : evar_map -> sorts * evar_map
+val is_sort_variable : evar_map -> sorts -> bool
+val whd_sort_variable : evar_map -> constr -> constr
+val set_leq_sort_variable : evar_map -> sorts -> sorts -> evar_map
+val define_sort_variable : evar_map -> sorts -> sorts -> evar_map
+
+(*********************************************************************)
+(* debug pretty-printer: *)
+
+val pr_evar_info : evar_info -> Pp.std_ppcmds
+val pr_evar_map : evar_map -> Pp.std_ppcmds
+val pr_evar_defs : evar_defs -> Pp.std_ppcmds
+val pr_sort_constraints : evar_map -> Pp.std_ppcmds
diff --git a/pretyping/indrec.ml b/pretyping/indrec.ml
index 0b9283ae..a587dd20 100644
--- a/pretyping/indrec.ml
+++ b/pretyping/indrec.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: indrec.ml,v 1.20.2.3 2004/07/16 19:30:44 herbelin Exp $ *)
+(* $Id: indrec.ml 7662 2005-12-17 22:03:35Z herbelin $ *)
open Pp
open Util
@@ -19,14 +19,21 @@ open Declarations
open Entries
open Inductive
open Inductiveops
-open Instantiate
open Environ
open Reductionops
open Typeops
open Type_errors
-open Indtypes (* pour les erreurs *)
open Safe_typing
open Nametab
+open Sign
+
+(* Errors related to recursors building *)
+type recursion_scheme_error =
+ | NotAllowedCaseAnalysis of bool * sorts * inductive
+ | BadInduction of bool * identifier * sorts
+ | NotMutualInScheme
+
+exception RecursionSchemeError of recursion_scheme_error
let make_prod_dep dep env = if dep then prod_name env else mkProd
let mkLambda_string s t c = mkLambda (Name (id_of_string s), t, c)
@@ -42,18 +49,18 @@ let mkLambda_string s t c = mkLambda (Name (id_of_string s), t, c)
lifter les paramètres globaux *)
let mis_make_case_com depopt env sigma (ind,mib,mip) kind =
- let lnamespar = mip.mind_params_ctxt in
+ let lnamespar = mib.mind_params_ctxt in
let dep = match depopt with
| None -> mip.mind_sort <> (Prop Null)
| Some d -> d
in
if not (List.exists ((=) kind) mip.mind_kelim) then
raise
- (InductiveError
+ (RecursionSchemeError
(NotAllowedCaseAnalysis
(dep,(new_sort_in_family kind),ind)));
- let nbargsprod = mip.mind_nrealargs + 1 in
+ let ndepar = mip.mind_nrealargs + 1 in
(* Pas génant car env ne sert pas à typer mais juste à renommer les Anonym *)
(* mais pas très joli ... (mais manque get_sort_of à ce niveau) *)
@@ -65,22 +72,28 @@ let mis_make_case_com depopt env sigma (ind,mib,mip) kind =
let rec add_branch env k =
if k = Array.length mip.mind_consnames then
let nbprod = k+1 in
- let indf = make_ind_family(ind,extended_rel_list nbprod lnamespar) in
- let lnamesar,_ = get_arity env indf in
+
+ let indf' = lift_inductive_family nbprod indf in
+ let arsign,_ = get_arity env indf' in
+ let depind = build_dependent_inductive env indf' in
+ let deparsign = (Anonymous,None,depind)::arsign in
+
let ci = make_default_case_info env RegularStyle ind in
- let depind = build_dependent_inductive env indf in
- let deparsign = (Anonymous,None,depind)::lnamesar in
- let p =
- it_mkLambda_or_LetIn_name env'
- (appvect
- (mkRel ((if dep then nbargsprod else mip.mind_nrealargs) + nbprod),
- if dep then extended_rel_vect 0 deparsign
- else extended_rel_vect 0 lnamesar))
- (if dep then deparsign else lnamesar) in
+ let pbody =
+ appvect
+ (mkRel (ndepar + nbprod),
+ if dep then extended_rel_vect 0 deparsign
+ else extended_rel_vect 1 arsign) in
+ let p =
+ it_mkLambda_or_LetIn_name env'
+ ((if dep then mkLambda_name env' else mkLambda)
+ (Anonymous,depind,pbody))
+ arsign
+ in
it_mkLambda_or_LetIn_name env'
- (mkCase (ci, lift nbargsprod p,
+ (mkCase (ci, lift ndepar p,
mkRel 1,
- rel_vect nbargsprod k))
+ rel_vect ndepar k))
deparsign
else
let cs = lift_constructor (k+1) constrs.(k) in
@@ -186,7 +199,7 @@ let type_rec_branch is_rec dep env sigma (vargs,depPvect,decP) tyi cs recargs =
let c = it_mkProd_or_LetIn base cs.cs_args in
process_constr env 0 c recargs nhyps []
-let make_rec_branch_arg env sigma (nparams,fvect,decF) f cstr recargs =
+let make_rec_branch_arg env sigma (nparrec,fvect,decF) f cstr recargs =
let process_pos env fk =
let rec prec env i hyps p =
let p',largs = whd_betadeltaiota_nolet_stack env sigma p in
@@ -198,7 +211,7 @@ let make_rec_branch_arg env sigma (nparams,fvect,decF) f cstr recargs =
let d = (n,Some b,t) in
mkLetIn (n,b,t,prec (push_rel d env) (i+1) (d::hyps) c)
| Ind _ ->
- let realargs = list_skipn nparams largs
+ let realargs = list_skipn nparrec largs
and arg = appvect (mkRel (i+1),extended_rel_vect 0 hyps) in
applist(lift i fk,realargs@[arg])
| _ -> assert false
@@ -239,10 +252,24 @@ let make_rec_branch_arg env sigma (nparams,fvect,decF) f cstr recargs =
in
process_constr env 0 f (List.rev cstr.cs_args, recargs)
+
+(* Cut a context ctx in 2 parts (ctx1,ctx2) with ctx1 containing k
+ variables *)
+let context_chop k ctx =
+ let rec chop_aux acc = function
+ | (0, l2) -> (List.rev acc, l2)
+ | (n, ((_,Some _,_ as h)::t)) -> chop_aux (h::acc) (n, t)
+ | (n, (h::t)) -> chop_aux (h::acc) (pred n, t)
+ | (_, []) -> failwith "context_chop"
+ in chop_aux [] (k,ctx)
+
+
(* Main function *)
-let mis_make_indrec env sigma listdepkind (ind,mib,mip) =
- let nparams = mip.mind_nparams in
- let lnamespar = mip.mind_params_ctxt in
+let mis_make_indrec env sigma listdepkind mib =
+ let nparams = mib.mind_nparams in
+ let nparrec = mib. mind_nparams_rec in
+ let lnonparrec,lnamesparrec =
+ context_chop (nparams-nparrec) mib.mind_params_ctxt in
let nrec = List.length listdepkind in
let depPvec =
Array.create mib.mind_ntypes (None : (bool * constr) option) in
@@ -257,6 +284,11 @@ let mis_make_indrec env sigma listdepkind (ind,mib,mip) =
assign nrec listdepkind in
let recargsvec =
Array.map (fun mip -> mip.mind_recargs) mib.mind_packets in
+ (* recarg information for non recursive parameters *)
+ let rec recargparn l n =
+ if n = 0 then l else recargparn (mk_norec::l) (n-1)
+ in
+ let recargpar = recargparn [] (nparams-nparrec) in
let make_one_rec p =
let makefix nbconstruct =
let rec mrec i ln ltyp ldef = function
@@ -267,59 +299,86 @@ let mis_make_indrec env sigma listdepkind (ind,mib,mip) =
(* arity in the context of the fixpoint, i.e.
P1..P_nrec f1..f_nbconstruct *)
- let args = extended_rel_list (nrec+nbconstruct) lnamespar in
+ let args = extended_rel_list (nrec+nbconstruct) lnamesparrec in
let indf = make_ind_family(indi,args) in
- let lnames,_ = get_arity env indf in
- let nar = mipi.mind_nrealargs in
- let dect = nar+nrec+nbconstruct in
+ let arsign,_ = get_arity env indf in
+ let depind = build_dependent_inductive env indf in
+ let deparsign = (Anonymous,None,depind)::arsign in
+
+ let nonrecpar = rel_context_length lnonparrec in
+ let larsign = rel_context_length deparsign in
+ let ndepar = larsign - nonrecpar in
+ let dect = larsign+nrec+nbconstruct in
- let branches =
(* constructors in context of the Cases expr, i.e.
- P1..P_nrec f1..f_nbconstruct F_1..F_nrec a_1..a_nar x:I *)
- let args' = extended_rel_list (dect+nrec+1) lnamespar in
- let indf' = make_ind_family(indi,args') in
+ P1..P_nrec f1..f_nbconstruct F_1..F_nrec a_1..a_nar x:I *)
+ let args' = extended_rel_list (dect+nrec) lnamesparrec in
+ let args'' = extended_rel_list ndepar lnonparrec in
+ let indf' = make_ind_family(indi,args'@args'') in
+
+ let branches =
let constrs = get_constructors env indf' in
- let vecfi = rel_vect (dect+1-i-nctyi) nctyi in
+ let fi = rel_vect (dect-i-nctyi) nctyi in
+ let vecfi = Array.map
+ (fun f -> appvect (f,extended_rel_vect ndepar lnonparrec))
+ fi
+ in
array_map3
- (make_rec_branch_arg env sigma (nparams,depPvec,nar+1))
- vecfi constrs (dest_subterms recargsvec.(tyi)) in
+ (make_rec_branch_arg env sigma
+ (nparrec,depPvec,larsign))
+ vecfi constrs (dest_subterms recargsvec.(tyi))
+ in
+
let j = (match depPvec.(tyi) with
| Some (_,c) when isRel c -> destRel c
- | _ -> assert false) in
+ | _ -> assert false)
+ in
+
+ (* Predicate in the context of the case *)
+
+ let depind' = build_dependent_inductive env indf' in
+ let arsign',_ = get_arity env indf' in
+ let deparsign' = (Anonymous,None,depind')::arsign' in
+
+ let pargs =
+ let nrpar = extended_rel_list (2*ndepar) lnonparrec
+ and nrar = if dep then extended_rel_list 0 deparsign'
+ else extended_rel_list 1 arsign'
+ in nrpar@nrar
+
+ in
+
+ (* body of i-th component of the mutual fixpoint *)
let deftyi =
let ci = make_default_case_info env RegularStyle indi in
- let indf' = lift_inductive_family nrec indf in
- let depind = build_dependent_inductive env indf' in
- let lnames' = Termops.lift_rel_context nrec lnames in
- let p =
- let arsign =
- if dep then (Anonymous,None,depind)::lnames' else lnames' in
- it_mkLambda_or_LetIn_name env
- (appvect
- (mkRel ((if dep then 1 else 0) + dect + j),
- extended_rel_vect 0 arsign)) arsign
+ let concl = applist (mkRel (dect+j+ndepar),pargs) in
+ let pred =
+ it_mkLambda_or_LetIn_name env
+ ((if dep then mkLambda_name env else mkLambda)
+ (Anonymous,depind',concl))
+ arsign'
in
it_mkLambda_or_LetIn_name env
- (lambda_create env
- (depind,mkCase (ci, lift (nar+1) p, mkRel 1, branches)))
- lnames'
+ (mkCase (ci, pred,
+ mkRel 1,
+ branches))
+ (lift_rel_context nrec deparsign)
in
- let typtyi =
- let ind = build_dependent_inductive env indf in
- it_mkProd_or_LetIn_name env
- (prod_create env
- (ind,
- (if dep then
- let ext_lnames = (Anonymous,None,ind)::lnames in
- let args = extended_rel_list 0 ext_lnames in
- applist (mkRel (nbconstruct+nar+j+1), args)
- else
- let args = extended_rel_list 1 lnames in
- applist (mkRel (nbconstruct+nar+j+1), args))))
- lnames
- in
- mrec (i+nctyi) (nar::ln) (typtyi::ltyp) (deftyi::ldef) rest
+
+ (* type of i-th component of the mutual fixpoint *)
+
+ let typtyi =
+ let concl =
+ let pargs = if dep then extended_rel_vect 0 deparsign
+ else extended_rel_vect 1 arsign
+ in appvect (mkRel (nbconstruct+ndepar+nonrecpar+j),pargs)
+ in it_mkProd_or_LetIn_name env
+ concl
+ deparsign
+ in
+ mrec (i+nctyi) (rel_context_nhyps arsign ::ln) (typtyi::ltyp)
+ (deftyi::ldef) rest
| [] ->
let fixn = Array.of_list (List.rev ln) in
let fixtyi = Array.of_list (List.rev ltyp) in
@@ -327,7 +386,7 @@ let mis_make_indrec env sigma listdepkind (ind,mib,mip) =
let names = Array.create nrec (Name(id_of_string "F")) in
mkFix ((fixn,p),(names,fixtyi,fixdef))
in
- mrec 0 [] [] []
+ mrec 0 [] [] []
in
let rec make_branch env i = function
| (indi,mibi,mipi,dep,_)::rest ->
@@ -338,8 +397,8 @@ let mis_make_indrec env sigma listdepkind (ind,mib,mip) =
make_branch env (i+j) rest
else
let recarg = (dest_subterms recargsvec.(tyi)).(j) in
- let vargs = extended_rel_list (nrec+i+j) lnamespar in
- let indf = (indi, vargs) in
+ let recarg = recargpar@recarg in
+ let vargs = extended_rel_list (nrec+i+j) lnamesparrec in
let cs = get_constructor (indi,mibi,mipi,vargs) (j+1) in
let p_0 =
type_rec_branch
@@ -353,23 +412,28 @@ let mis_make_indrec env sigma listdepkind (ind,mib,mip) =
in
let rec put_arity env i = function
| (indi,_,_,dep,kinds)::rest ->
- let indf = make_ind_family (indi,extended_rel_list i lnamespar) in
+ let indf = make_ind_family (indi,extended_rel_list i lnamesparrec) in
let typP = make_arity env dep indf (new_sort_in_family kinds) in
mkLambda_string "P" typP
(put_arity (push_rel (Anonymous,None,typP) env) (i+1) rest)
| [] ->
make_branch env 0 listdepkind
in
+
+ (* Body on make_one_rec *)
let (indi,mibi,mipi,dep,kind) = List.nth listdepkind p in
- let env' = push_rel_context lnamespar env in
+
if mis_is_recursive_subset
(List.map (fun (indi,_,_,_,_) -> snd indi) listdepkind)
mipi.mind_recargs
then
- it_mkLambda_or_LetIn_name env (put_arity env' 0 listdepkind) lnamespar
+ let env' = push_rel_context lnamesparrec env in
+ it_mkLambda_or_LetIn_name env (put_arity env' 0 listdepkind)
+ lnamesparrec
else
mis_make_case_com (Some dep) env sigma (indi,mibi,mipi) kind
in
+ (* Body of mis_make_indrec *)
list_tabulate make_one_rec nrec
(**********************************************************************)
@@ -385,20 +449,21 @@ let make_case_gen env = make_case_com None env
(**********************************************************************)
-(* [instanciate_indrec_scheme s rec] replace the sort of the scheme
+(* [instantiate_indrec_scheme s rec] replace the sort of the scheme
[rec] by [s] *)
let change_sort_arity sort =
let rec drec a = match kind_of_term a with
- | Cast (c,t) -> drec c
+ | Cast (c,_,_) -> drec c
| Prod (n,t,c) -> mkProd (n, t, drec c)
+ | LetIn (n,b,t,c) -> mkLetIn (n,b, t, drec c)
| Sort _ -> mkSort sort
| _ -> assert false
in
drec
(* [npar] is the number of expected arguments (then excluding letin's) *)
-let instanciate_indrec_scheme sort =
+let instantiate_indrec_scheme sort =
let rec drec npar elim =
match kind_of_term elim with
| Lambda (n,t,c) ->
@@ -407,13 +472,13 @@ let instanciate_indrec_scheme sort =
else
mkLambda (n, t, drec (npar-1) c)
| LetIn (n,b,t,c) -> mkLetIn (n,b,t,drec npar c)
- | _ -> anomaly "instanciate_indrec_scheme: wrong elimination type"
+ | _ -> anomaly "instantiate_indrec_scheme: wrong elimination type"
in
drec
(* Change the sort in the type of an inductive definition, builds the
corresponding eta-expanded term *)
-let instanciate_type_indrec_scheme sort npars term =
+let instantiate_type_indrec_scheme sort npars term =
let rec drec np elim =
match kind_of_term elim with
| Prod (n,t,c) ->
@@ -426,22 +491,27 @@ let instanciate_type_indrec_scheme sort npars term =
mkProd (n, t, c'), mkLambda (n, t, term')
| LetIn (n,b,t,c) -> let c',term' = drec np c in
mkLetIn (n,b,t,c'), mkLetIn (n,b,t,term')
- | _ -> anomaly "instanciate_type_indrec_scheme: wrong elimination type"
+ | _ -> anomaly "instantiate_type_indrec_scheme: wrong elimination type"
in
drec npars
(**********************************************************************)
(* Interface to build complex Scheme *)
+(* Check inductive types only occurs once
+(otherwise we obtain a meaning less scheme) *)
let check_arities listdepkind =
- List.iter
- (function (indi,mibi,mipi,dep,kind) ->
+ let _ = List.fold_left
+ (fun ln ((_,ni),mibi,mipi,dep,kind) ->
let id = mipi.mind_typename in
let kelim = mipi.mind_kelim in
- if not (List.exists ((=) kind) kelim) then
- raise
- (InductiveError (BadInduction (dep, id, new_sort_in_family kind))))
- listdepkind
+ if not (List.exists ((=) kind) kelim) then raise
+ (RecursionSchemeError (BadInduction (dep,id,new_sort_in_family kind)))
+ else if List.mem ni ln then raise
+ (RecursionSchemeError NotMutualInScheme)
+ else ni::ln)
+ [] listdepkind
+ in true
let build_mutual_indrec env sigma = function
| (mind,mib,mip,dep,s)::lrecspec ->
@@ -455,18 +525,18 @@ let build_mutual_indrec env sigma = function
let (mibi',mipi') = lookup_mind_specif env mind' in
(mind',mibi',mipi',dep',s')
else
- raise (InductiveError NotMutualInScheme))
+ raise (RecursionSchemeError NotMutualInScheme))
lrecspec)
in
let _ = check_arities listdepkind in
- mis_make_indrec env sigma listdepkind (mind,mib,mip)
+ mis_make_indrec env sigma listdepkind mib
| _ -> anomaly "build_indrec expects a non empty list of inductive types"
let build_indrec env sigma ind =
let (mib,mip) = lookup_mind_specif env ind in
let kind = family_of_sort mip.mind_sort in
let dep = kind <> InProp in
- List.hd (mis_make_indrec env sigma [(ind,mib,mip,dep,kind)] (ind,mib,mip))
+ List.hd (mis_make_indrec env sigma [(ind,mib,mip,dep,kind)] mib)
(**********************************************************************)
(* To handle old Case/Match syntax in Pretyping *)
@@ -483,7 +553,6 @@ let type_rec_branches recursive env sigma indt p c =
let tyi = snd ind in
let init_depPvec i = if i = tyi then Some(true,p) else None in
let depPvec = Array.init mib.mind_ntypes init_depPvec in
- let vargs = Array.of_list params in
let constructors = get_constructors env indf in
let lft =
array_map2
@@ -513,14 +582,14 @@ let lookup_eliminator ind_sp s =
let id = add_suffix ind_id (elimination_suffix s) in
(* Try first to get an eliminator defined in the same section as the *)
(* inductive type *)
- let ref = ConstRef (make_kn mp dp (label_of_id id)) in
+ let ref = ConstRef (make_con mp dp (label_of_id id)) in
try
let _ = sp_of_global ref in
- constr_of_reference ref
+ constr_of_global ref
with Not_found ->
(* Then try to get a user-defined eliminator in some other places *)
(* using short name (e.g. for "eq_rec") *)
- try constr_of_reference (Nametab.locate (make_short_qualid id))
+ try constr_of_global (Nametab.locate (make_short_qualid id))
with Not_found ->
errorlabstrm "default_elim"
(str "Cannot find the elimination combinator " ++
@@ -541,7 +610,7 @@ let lookup_eliminator ind_sp s =
with Not_found ->
(* Then try to get a user-defined eliminator in some other places *)
(* using short name (e.g. for "eq_rec") *)
- try constr_of_reference (Nametab.locate (make_short_qualid id))
+ try constr_of_global (Nametab.locate (make_short_qualid id))
with Not_found ->
errorlabstrm "default_elim"
(str "Cannot find the elimination combinator " ++
diff --git a/pretyping/indrec.mli b/pretyping/indrec.mli
index f6f76706..e5eb07f5 100644
--- a/pretyping/indrec.mli
+++ b/pretyping/indrec.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: indrec.mli,v 1.6.2.1 2004/07/16 19:30:45 herbelin Exp $ i*)
+(*i $Id: indrec.mli 7660 2005-12-17 21:13:48Z herbelin $ i*)
(*i*)
open Names
@@ -17,7 +17,16 @@ open Environ
open Evd
(*i*)
-(* Eliminations. *)
+(* Errors related to recursors building *)
+
+type recursion_scheme_error =
+ | NotAllowedCaseAnalysis of bool * sorts * inductive
+ | BadInduction of bool * identifier * sorts
+ | NotMutualInScheme
+
+exception RecursionSchemeError of recursion_scheme_error
+
+(** Eliminations *)
(* These functions build elimination predicate for Case tactic *)
@@ -29,11 +38,11 @@ val make_case_gen : env -> evar_map -> inductive -> sorts_family -> constr
of the inductive) *)
val build_indrec : env -> evar_map -> inductive -> constr
-val instanciate_indrec_scheme : sorts -> int -> constr -> constr
-val instanciate_type_indrec_scheme : sorts -> int -> constr -> types ->
+val instantiate_indrec_scheme : sorts -> int -> constr -> constr
+val instantiate_type_indrec_scheme : sorts -> int -> constr -> types ->
constr * types
-(* This builds complex [Scheme] *)
+(** Complex recursion schemes [Scheme] *)
val build_mutual_indrec :
env -> evar_map ->
@@ -41,7 +50,7 @@ val build_mutual_indrec :
* bool * sorts_family) list
-> constr list
-(* These are for old Case/Match typing *)
+(** Old Case/Match typing *)
val type_rec_branches : bool -> env -> evar_map -> inductive_type
-> constr -> constr -> constr array * constr
@@ -50,7 +59,11 @@ val make_rec_branch_arg :
int * ('b * constr) option array * int ->
constr -> constructor_summary -> wf_paths list -> constr
-(* *)
+(** Recursor names utilities *)
+
val lookup_eliminator : inductive -> sorts_family -> constr
val elimination_suffix : sorts_family -> string
val make_elimination_ident : identifier -> sorts_family -> identifier
+
+
+
diff --git a/pretyping/inductiveops.ml b/pretyping/inductiveops.ml
index c33a261b..57d966f1 100644
--- a/pretyping/inductiveops.ml
+++ b/pretyping/inductiveops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: inductiveops.ml,v 1.14.2.2 2004/12/29 12:15:00 herbelin Exp $ *)
+(* $Id: inductiveops.ml 8653 2006-03-22 09:41:17Z herbelin $ *)
open Util
open Names
@@ -18,6 +18,24 @@ open Declarations
open Environ
open Reductionops
+(* The following three functions are similar to the ones defined in
+ Inductive, but they expect an env *)
+
+let type_of_inductive env ind =
+ let specif = Inductive.lookup_mind_specif env ind in
+ Inductive.type_of_inductive specif
+
+(* Return type as quoted by the user *)
+let type_of_constructor env cstr =
+ let specif =
+ Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in
+ Inductive.type_of_constructor cstr specif
+
+(* Return constructor types in normal form *)
+let arities_of_constructors env ind =
+ let specif = Inductive.lookup_mind_specif env ind in
+ Inductive.arities_of_constructors ind specif
+
(* [inductive_family] = [inductive_instance] applied to global parameters *)
type inductive_family = inductive * constr list
@@ -73,6 +91,7 @@ let mis_nf_constructor_type (ind,mib,mip) j =
substl (list_tabulate make_Ik ntypes) specif.(j-1)
(* Arity of constructors excluding parameters and local defs *)
+
let mis_constr_nargs indsp =
let (mib,mip) = Global.lookup_inductive indsp in
let recargs = dest_subterms mip.mind_recargs in
@@ -87,7 +106,21 @@ let mis_constr_nargs_env env (kn,i) =
let mis_constructor_nargs_env env ((kn,i),j) =
let mib = Environ.lookup_mind kn env in
let mip = mib.mind_packets.(i) in
- recarg_length mip.mind_recargs j + mip.mind_nparams
+ recarg_length mip.mind_recargs j + mib.mind_nparams
+
+let constructor_nrealargs env (ind,j) =
+ let (_,mip) = Inductive.lookup_mind_specif env ind in
+ recarg_length mip.mind_recargs j
+
+let constructor_nrealhyps env (ind,j) =
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ mip.mind_consnrealdecls.(j-1)
+
+(* Length of arity (w/o local defs) *)
+
+let inductive_nargs env ind =
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
+ mip.mind_nrealargs + rel_context_nhyps mib.mind_params_ctxt
(* Annotation for cases *)
let make_case_info env ind style pats_source =
@@ -97,7 +130,8 @@ let make_case_info env ind style pats_source =
style = style;
source = pats_source } in
{ ci_ind = ind;
- ci_npar = mip.mind_nparams;
+ ci_npar = mib.mind_nparams;
+ ci_cstr_nargs = mip.mind_consnrealdecls;
ci_pp_info = print_info }
let make_default_case_info env style ind =
@@ -122,6 +156,7 @@ let lift_constructor n cs = {
cs_args = lift_rel_context n cs.cs_args;
cs_concl_realargs = Array.map (liftn n (cs.cs_nargs+1)) cs.cs_concl_realargs
}
+(* Accept less parameters than in the signature *)
let instantiate_params t args sign =
let rec inst s t = function
@@ -133,17 +168,17 @@ let instantiate_params t args sign =
(match kind_of_term t with
| LetIn(_,_,_,t) -> inst ((substl s b)::s) t (ctxt,args)
| _ -> anomaly"instantiate_params: type, ctxt and args mismatch")
- | [], [] -> substl s t
+ | _, [] -> substl s t
| _ -> anomaly"instantiate_params: type, ctxt and args mismatch"
in inst [] t (List.rev sign,args)
let get_constructor (ind,mib,mip,params) j =
assert (j <= Array.length mip.mind_consnames);
let typi = mis_nf_constructor_type (ind,mib,mip) j in
- let typi = instantiate_params typi params mip.mind_params_ctxt in
+ let typi = instantiate_params typi params mib.mind_params_ctxt in
let (args,ccl) = decompose_prod_assum typi in
let (_,allargs) = decompose_app ccl in
- let vargs = list_skipn mip.mind_nparams allargs in
+ let vargs = list_skipn (List.length params) allargs in
{ cs_cstr = ith_constructor_of_inductive ind j;
cs_params = params;
cs_nargs = rel_context_length args;
@@ -175,8 +210,7 @@ let build_dependent_constructor cs =
let build_dependent_inductive env ((ind, params) as indf) =
let arsign,_ = get_arity env indf in
- let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let nrealargs = mip.mind_nrealargs in
+ let nrealargs = List.length arsign in
applist
(mkInd ind,
(List.map (lift nrealargs) params)@(extended_rel_list 0 arsign))
@@ -189,7 +223,7 @@ let make_arity_signature env dep indf =
(* We need names everywhere *)
name_context env
((Anonymous,None,build_dependent_inductive env indf)::arsign)
- (* Costly: would be better to name one for all at definition time *)
+ (* Costly: would be better to name once for all at definition time *)
else
(* No need to enforce names *)
arsign
@@ -225,7 +259,7 @@ let find_rectype env sigma c =
match kind_of_term t with
| Ind ind ->
let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let (par,rargs) = list_chop mip.mind_nparams l in
+ let (par,rargs) = list_chop mib.mind_nparams l in
IndType((ind, par),rargs)
| _ -> raise Not_found
@@ -247,59 +281,60 @@ let find_coinductive env sigma c =
(***********************************************)
-(* find appropriate names for pattern variables. Useful in the
- Case tactic. *)
+(* find appropriate names for pattern variables. Useful in the Case
+ and Inversion (case_then_using et case_nodep_then_using) tactics. *)
-let is_dep_predicate env kelim pred nodep_ar =
- let rec srec env pval pt nodep_ar =
- let pt' = whd_betadeltaiota env Evd.empty pt in
+let is_predicate_explicitly_dep env pred nodep_ar =
+ let rec srec env pval nodep_ar =
let pv' = whd_betadeltaiota env Evd.empty pval in
- match kind_of_term pv', kind_of_term pt', kind_of_term nodep_ar with
- | Lambda (na,t,b), Prod (_,_,a), Prod (_,_,a') ->
- srec (push_rel_assum (na,t) env) b a a'
- | _, Prod (na,t,a), Prod (_,_,a') ->
- srec (push_rel_assum (na,t) env) (lift 1 pv') a a'
- | Lambda (_,_,b), Prod (_,_,_), _ -> (*dependent (mkRel 1) b*) true
- | _, Prod (_,_,_), _ -> true
- | _ -> false in
- srec env pred.uj_val pred.uj_type nodep_ar
-
-let is_dependent_elimination_predicate env pred indf =
- let (ind,params) = indf in
- let (_,mip) = Inductive.lookup_mind_specif env ind in
- let kelim = mip.mind_kelim in
- let arsign,s = get_arity env indf in
- let glob_t = it_mkProd_or_LetIn (mkSort s) arsign in
- is_dep_predicate env kelim pred glob_t
-
-let is_dep_arity env kelim predty nodep_ar =
- let rec srec pt nodep_ar =
- let pt' = whd_betadeltaiota env Evd.empty pt in
- match kind_of_term pt', kind_of_term nodep_ar with
- | Prod (_,a1,a2), Prod (_,a1',a2') -> srec a2 a2'
- | Prod (_,a1,a2), _ -> true
- | _ -> false in
- srec predty nodep_ar
-
-let is_dependent_elimination env predty indf =
- let (ind,params) = indf in
- let (_,mip) = Inductive.lookup_mind_specif env ind in
- let kelim = mip.mind_kelim in
+ match kind_of_term pv', kind_of_term nodep_ar with
+ | Lambda (na,t,b), Prod (_,_,a') ->
+ srec (push_rel_assum (na,t) env) b a'
+ | Lambda (na,_,_), _ ->
+
+ (* The following code has impact on the introduction names
+ given by the tactics "case" and "inversion": when the
+ elimination is not dependent, "case" uses Anonymous for
+ inductive types in Prop and names created by mkProd_name for
+ inductive types in Set/Type while "inversion" uses anonymous
+ for inductive types both in Prop and Set/Type !!
+
+ Previously, whether names were created or not relied on
+ whether the predicate created in Indrec.make_case_com had a
+ dependent arity or not. To avoid different predicates
+ printed the same in v8, all predicates built in indrec.ml
+ got a dependent arity (Aug 2004). The new way to decide
+ whether names have to be created or not is to use an
+ Anonymous or Named variable to enforce the expected
+ dependency status (of course, Anonymous implies non
+ dependent, but not conversely).
+
+ At the end, this is only to preserve the compatibility: a
+ check whether the predicate is actually dependent or not
+ would indeed be more natural! *)
+
+ na <> Anonymous
+
+ | _ -> anomaly "Non eta-expanded dep-expanded \"match\" predicate"
+ in
+ srec env pred nodep_ar
+
+let is_elim_predicate_explicitly_dependent env pred indf =
let arsign,s = get_arity env indf in
let glob_t = it_mkProd_or_LetIn (mkSort s) arsign in
- is_dep_arity env kelim predty glob_t
+ is_predicate_explicitly_dep env pred glob_t
let set_names env n brty =
let (ctxt,cl) = decompose_prod_n_assum n brty in
it_mkProd_or_LetIn_name env cl ctxt
let set_pattern_names env ind brv =
- let (_,mip) = Inductive.lookup_mind_specif env ind in
+ let (mib,mip) = Inductive.lookup_mind_specif env ind in
let arities =
Array.map
(fun c ->
rel_context_length (fst (decompose_prod_assum c)) -
- mip.mind_nparams)
+ mib.mind_nparams)
mip.mind_nf_lc in
array_map2 (set_names env) arities brv
@@ -308,8 +343,8 @@ let type_case_branches_with_names env indspec pj c =
let (ind,args) = indspec in
let (lbrty,conclty,_) = Inductive.type_case_branches env indspec pj c in
let (mib,mip) = Inductive.lookup_mind_specif env ind in
- let params = list_firstn mip.mind_nparams args in
- if is_dependent_elimination_predicate env pj (ind,params) then
+ let params = list_firstn mib.mind_nparams args in
+ if is_elim_predicate_explicitly_dependent env pj.uj_val (ind,params) then
(set_pattern_names env ind lbrty, conclty)
else (lbrty, conclty)
@@ -342,11 +377,15 @@ let control_only_guard env =
Array.iter control_rec tys;
Array.iter control_rec bds;
| Case(_,p,c,b) -> control_rec p;control_rec c;Array.iter control_rec b
- | Evar (_,cl) -> Array.iter control_rec cl
+ | Evar (_,cl) -> Array.iter control_rec cl
| App (c,cl) -> control_rec c; Array.iter control_rec cl
- | Cast (c1,c2) -> control_rec c1; control_rec c2
+ | Cast (c1,_, c2) -> control_rec c1; control_rec c2
| Prod (_,c1,c2) -> control_rec c1; control_rec c2
| Lambda (_,c1,c2) -> control_rec c1; control_rec c2
| LetIn (_,c1,c2,c3) -> control_rec c1; control_rec c2; control_rec c3
in
control_rec
+
+let subst_inductive subst (kn,i as ind) =
+ let kn' = Mod_subst.subst_kn subst kn in
+ if kn == kn' then ind else (kn',i)
diff --git a/pretyping/inductiveops.mli b/pretyping/inductiveops.mli
index 8cfa9b3c..2993eed3 100644
--- a/pretyping/inductiveops.mli
+++ b/pretyping/inductiveops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: inductiveops.mli,v 1.10.2.3 2005/01/21 17:19:37 herbelin Exp $ i*)
+(*i $Id: inductiveops.mli 7955 2006-01-30 22:56:15Z herbelin $ i*)
open Names
open Term
@@ -14,6 +14,17 @@ open Declarations
open Environ
open Evd
+(* The following three functions are similar to the ones defined in
+ Inductive, but they expect an env *)
+
+val type_of_inductive : env -> inductive -> types
+
+(* Return type as quoted by the user *)
+val type_of_constructor : env -> constructor -> types
+
+(* Return constructor types in normal form *)
+val arities_of_constructors : env -> inductive -> types array
+
(* An inductive type with its parameters *)
type inductive_family
val make_ind_family : inductive * constr list -> inductive_family
@@ -46,6 +57,11 @@ val mis_constr_nargs_env : env -> inductive -> int array
val mis_constructor_nargs_env : env -> constructor -> int
+val constructor_nrealargs : env -> constructor -> int
+val constructor_nrealhyps : env -> constructor -> int
+
+val inductive_nargs : env -> inductive -> int
+
type constructor_summary = {
cs_cstr : constructor;
cs_params : constr list;
@@ -74,9 +90,6 @@ val find_inductive : env -> evar_map -> constr -> inductive * constr list
val find_coinductive : env -> evar_map -> constr -> inductive * constr list
(********************)
-(* Determines if a case predicate type corresponds to dependent elimination *)
-val is_dependent_elimination :
- env -> types -> inductive_family -> bool
(* Builds the case predicate arity (dependent or not) *)
val arity_of_case_predicate :
@@ -91,3 +104,5 @@ val make_default_case_info : env -> case_style -> inductive -> case_info
(********************)
val control_only_guard : env -> types -> unit
+
+val subst_inductive : Mod_subst.substitution -> inductive -> inductive
diff --git a/pretyping/instantiate.ml b/pretyping/instantiate.ml
deleted file mode 100644
index 702cdfea..00000000
--- a/pretyping/instantiate.ml
+++ /dev/null
@@ -1,68 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(* $Id: instantiate.ml,v 1.3.2.1 2004/07/16 19:30:45 herbelin Exp $ *)
-
-open Pp
-open Util
-open Names
-open Term
-open Sign
-open Evd
-open Declarations
-open Environ
-
-let is_id_inst inst =
- let is_id (id,c) = match kind_of_term c with
- | Var id' -> id = id'
- | _ -> false
- in
- List.for_all is_id inst
-
-(* Vérifier que les instances des let-in sont compatibles ?? *)
-let instantiate_sign_including_let sign args =
- let rec instrec = function
- | ((id,b,_) :: sign, c::args) -> (id,c) :: (instrec (sign,args))
- | ([],[]) -> []
- | ([],_) | (_,[]) ->
- anomaly "Signature and its instance do not match"
- in
- instrec (sign,args)
-
-let instantiate_evar sign c args =
- let inst = instantiate_sign_including_let sign args in
- if is_id_inst inst then
- c
- else
- replace_vars inst c
-
-(* Existentials. *)
-
-let existential_type sigma (n,args) =
- let info =
- try Evd.map sigma n
- with Not_found ->
- anomaly ("Evar "^(string_of_existential n)^" was not declared") in
- let hyps = info.evar_hyps in
- instantiate_evar hyps info.evar_concl (Array.to_list args)
-
-exception NotInstantiatedEvar
-
-let existential_value sigma (n,args) =
- let info = Evd.map sigma n in
- let hyps = info.evar_hyps in
- match evar_body info with
- | Evar_defined c ->
- instantiate_evar hyps c (Array.to_list args)
- | Evar_empty ->
- raise NotInstantiatedEvar
-
-let existential_opt_value sigma ev =
- try Some (existential_value sigma ev)
- with NotInstantiatedEvar -> None
-
diff --git a/pretyping/instantiate.mli b/pretyping/instantiate.mli
deleted file mode 100644
index 44c4d579..00000000
--- a/pretyping/instantiate.mli
+++ /dev/null
@@ -1,25 +0,0 @@
-(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
-(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
-(************************************************************************)
-
-(*i $Id: instantiate.mli,v 1.2.14.1 2004/07/16 19:30:45 herbelin Exp $ i*)
-
-(*i*)
-open Names
-open Term
-open Evd
-open Sign
-open Environ
-(*i*)
-
-(*s [existential_value sigma ev] raises [NotInstantiatedEvar] if [ev] has
-no body and [Not_found] if it does not exist in [sigma] *)
-
-exception NotInstantiatedEvar
-val existential_value : evar_map -> existential -> constr
-val existential_type : evar_map -> existential -> types
-val existential_opt_value : evar_map -> existential -> constr option
diff --git a/pretyping/matching.ml b/pretyping/matching.ml
index bdab3b5b..5ee245b5 100644
--- a/pretyping/matching.ml
+++ b/pretyping/matching.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: matching.ml,v 1.3.2.1 2004/07/16 19:30:45 herbelin Exp $ *)
+(* $Id: matching.ml 7970 2006-02-01 15:09:07Z herbelin $ *)
(*i*)
open Util
@@ -89,11 +89,15 @@ let matches_core convert allow_partial_app pat c =
| PMeta (Some n), m ->
let depth = List.length stk in
- let frels = Intset.elements (free_rels cT) in
- if List.for_all (fun i -> i > depth) frels then
- constrain (n,lift (-depth) cT) sigma
- else
- raise PatternMatchingFailure
+ if depth = 0 then
+ (* Optimisation *)
+ constrain (n,cT) sigma
+ else
+ let frels = Intset.elements (free_rels cT) in
+ if List.for_all (fun i -> i > depth) frels then
+ constrain (n,lift (-depth) cT) sigma
+ else
+ raise PatternMatchingFailure
| PMeta None, m -> sigma
@@ -101,7 +105,7 @@ let matches_core convert allow_partial_app pat c =
| PVar v1, Var v2 when v1 = v2 -> sigma
- | PRef ref, _ when constr_of_reference ref = cT -> sigma
+ | PRef ref, _ when constr_of_global ref = cT -> sigma
| PRel n1, Rel n2 when n1 = n2 -> sigma
@@ -109,6 +113,9 @@ let matches_core convert allow_partial_app pat c =
| PSort (RType _), Sort (Type _) -> sigma
+ | PApp (PApp (h, a1), a2), _ ->
+ sorec stk sigma (PApp(h,Array.append a1 a2)) t
+
| PApp (PMeta (Some n),args1), App (c2,args2) when allow_partial_app ->
let p = Array.length args2 - Array.length args1 in
if p>=0 then
@@ -139,7 +146,7 @@ let matches_core convert allow_partial_app pat c =
| PRef (ConstRef _ as ref), _ when convert <> None ->
let (env,evars) = out_some convert in
- let c = constr_of_reference ref in
+ let c = constr_of_global ref in
if is_conv env evars c cT then sigma
else raise PatternMatchingFailure
@@ -176,15 +183,15 @@ let special_meta = (-1)
(* Tries to match a subterm of [c] with [pat] *)
let rec sub_match nocc pat c =
match kind_of_term c with
- | Cast (c1,c2) ->
+ | Cast (c1,k,c2) ->
(try authorized_occ nocc ((matches pat c), mkMeta special_meta) with
| PatternMatchingFailure ->
let (lm,lc) = try_sub_match nocc pat [c1] in
- (lm,mkCast (List.hd lc, c2))
+ (lm,mkCast (List.hd lc, k,c2))
| NextOccurrence nocc ->
let (lm,lc) = try_sub_match (nocc - 1) pat [c1] in
- (lm,mkCast (List.hd lc, c2)))
- | Lambda (x,c1,c2) ->
+ (lm,mkCast (List.hd lc, k,c2)))
+ | Lambda (x,c1,c2) ->
(try authorized_occ nocc ((matches pat c), mkMeta special_meta) with
| PatternMatchingFailure ->
let (lm,lc) = try_sub_match nocc pat [c1;c2] in
@@ -242,6 +249,10 @@ and try_sub_match nocc pat lc =
| NextOccurrence nocc -> try_sub_match_rec nocc pat (lacc@[c]) tl) in
try_sub_match_rec nocc pat [] lc
+let match_subterm nocc pat c =
+ try sub_match nocc pat c
+ with NextOccurrence _ -> raise PatternMatchingFailure
+
let is_matching pat n =
try let _ = matches pat n in true
with PatternMatchingFailure -> false
diff --git a/pretyping/matching.mli b/pretyping/matching.mli
index 2f666880..e6065c68 100644
--- a/pretyping/matching.mli
+++ b/pretyping/matching.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: matching.mli,v 1.3.2.2 2005/01/21 16:42:37 herbelin Exp $ i*)
+(*i $Id: matching.mli 6616 2005-01-21 17:18:23Z herbelin $ i*)
(*i*)
open Names
@@ -39,11 +39,10 @@ val is_matching : constr_pattern -> constr -> bool
val matches_conv :env -> Evd.evar_map -> constr_pattern -> constr -> patvar_map
-(* To skip to the next occurrence *)
-exception NextOccurrence of int
-
-(* Tries to match a **closed** subterm of [c] with [pat] *)
-val sub_match : int -> constr_pattern -> constr -> patvar_map * constr
+(* [match_subterm n pat c] returns the substitution and the context
+ corresponding to the [n+1]th **closed** subterm of [c] matching [pat];
+ It raises PatternMatchingFailure if no such matching exists *)
+val match_subterm : int -> constr_pattern -> constr -> patvar_map * constr
(* [is_matching_conv env sigma pat c] tells if [c] matches against [pat]
up to conversion for constants in patterns *)
diff --git a/pretyping/pattern.ml b/pretyping/pattern.ml
index f58a12c6..390b884c 100644
--- a/pretyping/pattern.ml
+++ b/pretyping/pattern.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pattern.ml,v 1.24.2.2 2004/11/26 17:51:52 herbelin Exp $ *)
+(* $Id: pattern.ml 7732 2005-12-26 13:51:24Z herbelin $ *)
open Util
open Names
@@ -17,18 +17,13 @@ open Rawterm
open Environ
open Nametab
open Pp
+open Mod_subst
(* Metavariables *)
type patvar_map = (patvar * constr) list
-let patvar_of_int n =
- let p = if !Options.v7 & not (Options.do_translate ()) then "?" else "X"
- in
- Names.id_of_string (p ^ string_of_int n)
let pr_patvar = pr_id
-let patvar_of_int_v7 n = Names.id_of_string ("?" ^ string_of_int n)
-
(* Patterns *)
type constr_pattern =
@@ -62,57 +57,6 @@ let rec occur_meta_pattern = function
| PMeta _ | PSoApp _ -> true
| PEvar _ | PVar _ | PRef _ | PRel _ | PSort _ | PFix _ | PCoFix _ -> false
-let rec subst_pattern subst pat = match pat with
- | PRef ref ->
- let ref' = subst_global subst ref in
- if ref' == ref then pat else
- PRef ref'
- | PVar _
- | PEvar _
- | PRel _ -> pat
- | PApp (f,args) ->
- let f' = subst_pattern subst f in
- let args' = array_smartmap (subst_pattern subst) args in
- if f' == f && args' == args then pat else
- PApp (f',args')
- | PSoApp (i,args) ->
- let args' = list_smartmap (subst_pattern subst) args in
- if args' == args then pat else
- PSoApp (i,args')
- | PLambda (name,c1,c2) ->
- let c1' = subst_pattern subst c1 in
- let c2' = subst_pattern subst c2 in
- if c1' == c1 && c2' == c2 then pat else
- PLambda (name,c1',c2')
- | PProd (name,c1,c2) ->
- let c1' = subst_pattern subst c1 in
- let c2' = subst_pattern subst c2 in
- if c1' == c1 && c2' == c2 then pat else
- PProd (name,c1',c2')
- | PLetIn (name,c1,c2) ->
- let c1' = subst_pattern subst c1 in
- let c2' = subst_pattern subst c2 in
- if c1' == c1 && c2' == c2 then pat else
- PLetIn (name,c1',c2')
- | PSort _
- | PMeta _ -> pat
- | PCase (cs,typ, c, branches) ->
- let typ' = option_smartmap (subst_pattern subst) typ in
- let c' = subst_pattern subst c in
- let branches' = array_smartmap (subst_pattern subst) branches in
- if typ' == typ && c' == c && branches' == branches then pat else
- PCase(cs,typ', c', branches')
- | PFix fixpoint ->
- let cstr = mkFix fixpoint in
- let fixpoint' = destFix (subst_mps subst cstr) in
- if fixpoint' == fixpoint then pat else
- PFix fixpoint'
- | PCoFix cofixpoint ->
- let cstr = mkCoFix cofixpoint in
- let cofixpoint' = destCoFix (subst_mps subst cstr) in
- if cofixpoint' == cofixpoint then pat else
- PCoFix cofixpoint'
-
type constr_label =
| ConstNode of constant
| IndNode of inductive
@@ -121,33 +65,14 @@ type constr_label =
exception BoundPattern;;
-let label_of_ref = function
- | ConstRef sp -> ConstNode sp
- | IndRef sp -> IndNode sp
- | ConstructRef sp -> CstrNode sp
- | VarRef id -> VarNode id
-
-let ref_of_label = function
- | ConstNode sp -> ConstRef sp
- | IndNode sp -> IndRef sp
- | CstrNode sp -> ConstructRef sp
- | VarNode id -> VarRef id
-
-let subst_label subst cstl =
- let ref = ref_of_label cstl in
- let ref' = subst_global subst ref in
- if ref' == ref then cstl else
- label_of_ref ref'
-
-
let rec head_pattern_bound t =
match t with
| PProd (_,_,b) -> head_pattern_bound b
| PLetIn (_,_,b) -> head_pattern_bound b
| PApp (c,args) -> head_pattern_bound c
| PCase (_,p,c,br) -> head_pattern_bound c
- | PRef r -> label_of_ref r
- | PVar id -> VarNode id
+ | PRef r -> r
+ | PVar id -> VarRef id
| PEvar _ | PRel _ | PMeta _ | PSoApp _ | PSort _ | PFix _
-> raise BoundPattern
(* Perhaps they were arguments, but we don't beta-reduce *)
@@ -155,10 +80,10 @@ let rec head_pattern_bound t =
| PCoFix _ -> anomaly "head_pattern_bound: not a type"
let head_of_constr_reference c = match kind_of_term c with
- | Const sp -> ConstNode sp
- | Construct sp -> CstrNode sp
- | Ind sp -> IndNode sp
- | Var id -> VarNode id
+ | Const sp -> ConstRef sp
+ | Construct sp -> ConstructRef sp
+ | Ind sp -> IndRef sp
+ | Var id -> VarRef id
| _ -> anomaly "Not a rigid reference"
let rec pattern_of_constr t =
@@ -168,7 +93,7 @@ let rec pattern_of_constr t =
| Var id -> PVar id
| Sort (Prop c) -> PSort (RProp c)
| Sort (Type _) -> PSort (RType None)
- | Cast (c,_) -> pattern_of_constr c
+ | Cast (c,_,_) -> pattern_of_constr c
| LetIn (na,c,_,b) -> PLetIn (na,pattern_of_constr c,pattern_of_constr b)
| Prod (na,c,b) -> PProd (na,pattern_of_constr c,pattern_of_constr b)
| Lambda (na,c,b) -> PLambda (na,pattern_of_constr c,pattern_of_constr b)
@@ -198,9 +123,61 @@ let rec inst lvar = function
(* Non recursive *)
| (PEvar _ | PRel _ | PRef _ | PSort _ | PMeta _ as x) -> x
(* Bound to terms *)
- | (PFix _ | PCoFix _ as r) ->
+ | (PFix _ | PCoFix _) ->
error ("Not instantiable pattern")
+let rec subst_pattern subst pat = match pat with
+ | PRef ref ->
+ let ref',t = subst_global subst ref in
+ if ref' == ref then pat else
+ pattern_of_constr t
+ | PVar _
+ | PEvar _
+ | PRel _ -> pat
+ | PApp (f,args) ->
+ let f' = subst_pattern subst f in
+ let args' = array_smartmap (subst_pattern subst) args in
+ if f' == f && args' == args then pat else
+ PApp (f',args')
+ | PSoApp (i,args) ->
+ let args' = list_smartmap (subst_pattern subst) args in
+ if args' == args then pat else
+ PSoApp (i,args')
+ | PLambda (name,c1,c2) ->
+ let c1' = subst_pattern subst c1 in
+ let c2' = subst_pattern subst c2 in
+ if c1' == c1 && c2' == c2 then pat else
+ PLambda (name,c1',c2')
+ | PProd (name,c1,c2) ->
+ let c1' = subst_pattern subst c1 in
+ let c2' = subst_pattern subst c2 in
+ if c1' == c1 && c2' == c2 then pat else
+ PProd (name,c1',c2')
+ | PLetIn (name,c1,c2) ->
+ let c1' = subst_pattern subst c1 in
+ let c2' = subst_pattern subst c2 in
+ if c1' == c1 && c2' == c2 then pat else
+ PLetIn (name,c1',c2')
+ | PSort _
+ | PMeta _ -> pat
+ | PCase (cs,typ, c, branches) ->
+ let typ' = option_smartmap (subst_pattern subst) typ in
+ let c' = subst_pattern subst c in
+ let branches' = array_smartmap (subst_pattern subst) branches in
+ if typ' == typ && c' == c && branches' == branches then pat else
+ PCase(cs,typ', c', branches')
+ | PFix fixpoint ->
+ let cstr = mkFix fixpoint in
+ let fixpoint' = destFix (subst_mps subst cstr) in
+ if fixpoint' == fixpoint then pat else
+ PFix fixpoint'
+ | PCoFix cofixpoint ->
+ let cstr = mkCoFix cofixpoint in
+ let cofixpoint' = destCoFix (subst_mps subst cstr) in
+ if cofixpoint' == cofixpoint then pat else
+ PCoFix cofixpoint'
+
+
let instantiate_pattern = inst
let rec pat_of_raw metas vars = function
@@ -230,30 +207,25 @@ let rec pat_of_raw metas vars = function
PSort s
| RHole _ ->
PMeta None
- | RCast (_,c,t) ->
+ | RCast (_,c,_,t) ->
Options.if_verbose
Pp.warning "Cast not taken into account in constr pattern";
pat_of_raw metas vars c
- | ROrderedCase (_,st,po,c,br,_) ->
- PCase ((None,st),option_app (pat_of_raw metas vars) po,
- pat_of_raw metas vars c,
- Array.map (pat_of_raw metas vars) br)
| RIf (_,c,(_,None),b1,b2) ->
PCase ((None,IfStyle),None, pat_of_raw metas vars c,
[|pat_of_raw metas vars b1; pat_of_raw metas vars b2|])
- | RCases (loc,(po,_),[c,_],brs) ->
+ | RCases (loc,None,[c,_],brs) ->
let sp =
match brs with
| (_,_,[PatCstr(_,(ind,_),_,_)],_)::_ -> Some ind
| _ -> None in
- (* When po disappears: switch to rtn type *)
- PCase ((sp,Term.RegularStyle),option_app (pat_of_raw metas vars) po,
+ PCase ((sp,Term.RegularStyle),None,
pat_of_raw metas vars c,
Array.init (List.length brs)
(pat_of_raw_branch loc metas vars sp brs))
| r ->
let loc = loc_of_rawconstr r in
- user_err_loc (loc,"pattern_of_rawconstr", Pp.str "Not supported pattern")
+ user_err_loc (loc,"pattern_of_rawconstr", Pp.str "Pattern not supported")
and pat_of_raw_branch loc metas vars ind brs i =
let bri = List.filter
diff --git a/pretyping/pattern.mli b/pretyping/pattern.mli
index cf0d4528..25a57ed2 100644
--- a/pretyping/pattern.mli
+++ b/pretyping/pattern.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pattern.mli,v 1.17.2.1 2004/07/16 19:30:45 herbelin Exp $ i*)
+(*i $Id: pattern.mli 7732 2005-12-26 13:51:24Z herbelin $ i*)
(*i*)
open Pp
@@ -17,6 +17,7 @@ open Environ
open Libnames
open Nametab
open Rawterm
+open Mod_subst
(*i*)
(* Pattern variables *)
@@ -24,10 +25,6 @@ open Rawterm
type patvar_map = (patvar * constr) list
val pr_patvar : patvar -> std_ppcmds
-(* Only for v7 parsing/printing *)
-val patvar_of_int : int -> patvar
-val patvar_of_int_v7 : int -> patvar
-
(* Patterns *)
type constr_pattern =
@@ -51,28 +48,18 @@ val occur_meta_pattern : constr_pattern -> bool
val subst_pattern : substitution -> constr_pattern -> constr_pattern
-type constr_label =
- | ConstNode of constant
- | IndNode of inductive
- | CstrNode of constructor
- | VarNode of identifier
-
-val label_of_ref : global_reference -> constr_label
-
-val subst_label : substitution -> constr_label -> constr_label
-
exception BoundPattern
(* [head_pattern_bound t] extracts the head variable/constant of the
type [t] or raises [BoundPattern] (even if a sort); it raises an anomaly
if [t] is an abstraction *)
-val head_pattern_bound : constr_pattern -> constr_label
+val head_pattern_bound : constr_pattern -> global_reference
(* [head_of_constr_reference c] assumes [r] denotes a reference and
returns its label; raises an anomaly otherwise *)
-val head_of_constr_reference : Term.constr -> constr_label
+val head_of_constr_reference : Term.constr -> global_reference
(* [pattern_of_constr c] translates a term [c] with metavariables into
a pattern; currently, no destructor (Cases, Fix, Cofix) and no
diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml
index fee1522f..48295c92 100644
--- a/pretyping/pretype_errors.ml
+++ b/pretyping/pretype_errors.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pretype_errors.ml,v 1.25.2.2 2004/07/16 19:30:45 herbelin Exp $ *)
+(* $Id: pretype_errors.ml 8688 2006-04-07 15:08:12Z msozeau $ *)
open Util
open Stdpp
@@ -24,8 +24,12 @@ type pretype_error =
| CantFindCaseType of constr
(* Unification *)
| OccurCheck of existential_key * constr
- | NotClean of existential_key * constr * hole_kind
- | UnsolvableImplicit of hole_kind
+ | NotClean of existential_key * constr * Evd.hole_kind
+ | UnsolvableImplicit of Evd.hole_kind
+ | CannotUnify of constr * constr
+ | CannotUnifyBindingType of constr * constr
+ | CannotGeneralize of constr
+ | NoOccurrenceFound of constr
(* Pretyping *)
| VarNotFound of identifier
| UnexpectedType of constr * constr
@@ -33,6 +37,12 @@ type pretype_error =
exception PretypeError of env * pretype_error
+let precatchable_exception = function
+ | Util.UserError _ | TypeError _ | PretypeError _
+ | Stdpp.Exc_located(_,(Util.UserError _ | TypeError _ |
+ Nametab.GlobalizationError _ | PretypeError _)) -> true
+ | _ -> false
+
let nf_evar = Reductionops.nf_evar
let j_nf_evar sigma j =
{ uj_val = nf_evar sigma j.uj_val;
@@ -43,7 +53,7 @@ let tj_nf_evar sigma {utj_val=v;utj_type=t} =
{utj_val=type_app (nf_evar sigma) v;utj_type=t}
let env_ise sigma env =
- let sign = named_context env in
+ let sign = named_context_val env in
let ctxt = rel_context env in
let env0 = reset_with_named_context sign env in
Sign.fold_rel_context
@@ -126,6 +136,9 @@ let error_ill_typed_rec_body_loc loc env sigma i na jl tys =
IllTypedRecBody (i,na,jv_nf_evar sigma jl,
Array.map (nf_evar sigma) tys))
+let error_not_a_type_loc loc env sigma j =
+ raise_located_type_error (loc, env, sigma, NotAType (j_nf_evar sigma j))
+
(*s Implicit arguments synthesis errors. It is hard to find
a precise location. *)
@@ -141,6 +154,12 @@ let error_not_clean env sigma ev c (loc,k) =
let error_unsolvable_implicit loc env sigma e =
raise_with_loc loc (PretypeError (env_ise sigma env, UnsolvableImplicit e))
+let error_cannot_unify env sigma (m,n) =
+ raise (PretypeError (env_ise sigma env,CannotUnify (m,n)))
+
+let error_cannot_coerce env sigma (m,n) =
+ raise (PretypeError (env_ise sigma env,CannotUnify (m,n)))
+
(*s Ml Case errors *)
let error_cant_find_case_type_loc loc env sigma expr =
diff --git a/pretyping/pretype_errors.mli b/pretyping/pretype_errors.mli
index ebeff99d..3c78d48d 100644
--- a/pretyping/pretype_errors.mli
+++ b/pretyping/pretype_errors.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pretype_errors.mli,v 1.25.2.3 2004/07/16 19:30:45 herbelin Exp $ i*)
+(*i $Id: pretype_errors.mli 8688 2006-04-07 15:08:12Z msozeau $ i*)
(*i*)
open Pp
@@ -26,8 +26,12 @@ type pretype_error =
| CantFindCaseType of constr
(* Unification *)
| OccurCheck of existential_key * constr
- | NotClean of existential_key * constr * hole_kind
- | UnsolvableImplicit of hole_kind
+ | NotClean of existential_key * constr * Evd.hole_kind
+ | UnsolvableImplicit of Evd.hole_kind
+ | CannotUnify of constr * constr
+ | CannotUnifyBindingType of constr * constr
+ | CannotGeneralize of constr
+ | NoOccurrenceFound of constr
(* Pretyping *)
| VarNotFound of identifier
| UnexpectedType of constr * constr
@@ -35,6 +39,8 @@ type pretype_error =
exception PretypeError of env * pretype_error
+val precatchable_exception : exn -> bool
+
(* Presenting terms without solved evars *)
val nf_evar : Evd.evar_map -> constr -> constr
val j_nf_evar : Evd.evar_map -> unsafe_judgment -> unsafe_judgment
@@ -73,14 +79,22 @@ val error_ill_typed_rec_body_loc :
loc -> env -> Evd.evar_map ->
int -> name array -> unsafe_judgment array -> types array -> 'b
+val error_not_a_type_loc :
+ loc -> env -> Evd.evar_map -> unsafe_judgment -> 'b
+
+val error_cannot_coerce : env -> Evd.evar_map -> constr * constr -> 'b
+
(*s Implicit arguments synthesis errors *)
val error_occur_check : env -> Evd.evar_map -> existential_key -> constr -> 'b
val error_not_clean :
- env -> Evd.evar_map -> existential_key -> constr -> loc * hole_kind -> 'b
+ env -> Evd.evar_map -> existential_key -> constr -> loc * Evd.hole_kind -> 'b
+
+val error_unsolvable_implicit :
+ loc -> env -> Evd.evar_map -> Evd.hole_kind -> 'b
-val error_unsolvable_implicit : loc -> env -> Evd.evar_map -> hole_kind -> 'b
+val error_cannot_unify : env -> Evd.evar_map -> constr * constr -> 'b
(*s Ml Case errors *)
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index bb0e74bb..2d1e297f 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pretyping.ml,v 1.123.2.5 2005/11/29 21:40:52 letouzey Exp $ *)
+(* $Id: pretyping.ml 8695 2006-04-10 16:33:52Z msozeau $ *)
open Pp
open Util
@@ -20,6 +20,7 @@ open Environ
open Type_errors
open Typeops
open Libnames
+open Nameops
open Classops
open List
open Recordops
@@ -27,86 +28,18 @@ open Evarutil
open Pretype_errors
open Rawterm
open Evarconv
-open Coercion
open Pattern
open Dyn
+type typing_constraint = OfType of types option | IsType
+type var_map = (identifier * unsafe_judgment) list
+type unbound_ltac_var_map = (identifier * identifier option) list
(************************************************************************)
(* This concerns Cases *)
open Declarations
open Inductive
open Inductiveops
-open Instantiate
-
-let lift_context n l =
- let k = List.length l in
- list_map_i (fun i (name,c) -> (name,liftn n (k-i) c)) 0 l
-
-let transform_rec loc env sigma (pj,c,lf) indt =
- let p = pj.uj_val in
- let (indf,realargs) = dest_ind_type indt in
- let (ind,params) = dest_ind_family indf in
- let (mib,mip) = lookup_mind_specif env ind in
- let recargs = mip.mind_recargs in
- let mI = mkInd ind in
- let ci = make_default_case_info env (if Options.do_translate() then RegularStyle else MatchStyle) ind in
- let nconstr = Array.length mip.mind_consnames in
- if Array.length lf <> nconstr then
- (let cj = {uj_val=c; uj_type=mkAppliedInd indt} in
- error_number_branches_loc loc env sigma cj nconstr);
- let tyi = snd ind in
- if mis_is_recursive_subset [tyi] recargs then
- let dep =
- is_dependent_elimination env (nf_evar sigma pj.uj_type) indf in
- let init_depFvec i = if i = tyi then Some(dep,mkRel 1) else None in
- let depFvec = Array.init mib.mind_ntypes init_depFvec in
- (* build now the fixpoint *)
- let lnames,_ = get_arity env indf in
- let nar = List.length lnames in
- let nparams = mip.mind_nparams in
- let constrs = get_constructors env (lift_inductive_family (nar+2) indf) in
- let branches =
- array_map3
- (fun f t reca ->
- whd_beta
- (Indrec.make_rec_branch_arg env sigma
- (nparams,depFvec,nar+1)
- f t reca))
- (Array.map (lift (nar+2)) lf) constrs (dest_subterms recargs)
- in
- let deffix =
- it_mkLambda_or_LetIn_name env
- (lambda_create env
- (applist (mI,List.append (List.map (lift (nar+1)) params)
- (extended_rel_list 0 lnames)),
- mkCase (ci, lift (nar+2) p, mkRel 1, branches)))
- (lift_rel_context 1 lnames)
- in
- if noccurn 1 deffix then
- whd_beta (applist (pop deffix,realargs@[c]))
- else
- let ind = applist (mI,(List.append
- (List.map (lift nar) params)
- (extended_rel_list 0 lnames))) in
- let typPfix =
- it_mkProd_or_LetIn_name env
- (prod_create env
- (ind,
- (if dep then
- let ext_lnames = (Anonymous,None,ind)::lnames in
- let args = extended_rel_list 0 ext_lnames in
- whd_beta (applist (lift (nar+1) p, args))
- else
- let args = extended_rel_list 1 lnames in
- whd_beta (applist (lift (nar+1) p, args)))))
- lnames in
- let fix =
- mkFix (([|nar|],0),
- ([|Name(id_of_string "F")|],[|typPfix|],[|deffix|])) in
- applist (fix,realargs@[c])
- else
- mkCase (ci, p, c, lf)
(************************************************************************)
@@ -114,909 +47,667 @@ let transform_rec loc env sigma (pj,c,lf) indt =
let ((constr_in : constr -> Dyn.t),
(constr_out : Dyn.t -> constr)) = create "constr"
-let mt_evd = Evd.empty
-
-let vect_lift_type = Array.mapi (fun i t -> type_app (lift i) t)
-
-(* Utilisé pour inférer le prédicat des Cases *)
-(* Semble exagérement fort *)
-(* Faudra préférer une unification entre les types de toutes les clauses *)
-(* et autoriser des ? à rester dans le résultat de l'unification *)
-
-let evar_type_fixpoint loc env isevars lna lar vdefj =
- let lt = Array.length vdefj in
- if Array.length lar = lt then
- for i = 0 to lt-1 do
- if not (the_conv_x_leq env isevars
- (vdefj.(i)).uj_type
- (lift lt lar.(i))) then
- error_ill_typed_rec_body_loc loc env (evars_of isevars)
- i lna vdefj lar
- done
-
-let check_branches_message loc env isevars c (explft,lft) =
- for i = 0 to Array.length explft - 1 do
- if not (the_conv_x_leq env isevars lft.(i) explft.(i)) then
- let sigma = evars_of isevars in
- error_ill_formed_branch_loc loc env sigma c i lft.(i) explft.(i)
- done
-
-(* coerce to tycon if any *)
-let inh_conv_coerce_to_tycon loc env isevars j = function
- | None -> j
- | Some typ -> inh_conv_coerce_to loc env isevars j typ
-
-let push_rels vars env = List.fold_right push_rel vars env
-
-(*
-let evar_type_case isevars env ct pt lft p c =
- let (mind,bty,rslty) = type_case_branches env (evars_of isevars) ct pt p c
- in check_branches_message isevars env (c,ct) (bty,lft); (mind,rslty)
-*)
-
-let strip_meta id = (* For Grammar v7 compatibility *)
- let s = string_of_id id in
- if s.[0]='$' then id_of_string (String.sub s 1 (String.length s - 1))
- else id
-
-let pretype_id loc env (lvar,unbndltacvars) id =
- let id = strip_meta id in (* May happen in tactics defined by Grammar *)
- try
- let (n,typ) = lookup_rel_id id (rel_context env) in
- { uj_val = mkRel n; uj_type = type_app (lift n) typ }
- with Not_found ->
- try
- List.assoc id lvar
- with Not_found ->
- try
- let (_,_,typ) = lookup_named id env in
- { uj_val = mkVar id; uj_type = typ }
- with Not_found ->
- try (* To build a nicer ltac error message *)
- match List.assoc id unbndltacvars with
- | None -> user_err_loc (loc,"",
- str (string_of_id id ^ " ist not bound to a term"))
- | Some id0 -> Pretype_errors.error_var_not_found_loc loc id0
- with Not_found ->
- error_var_not_found_loc loc id
-
-(* make a dependent predicate from an undependent one *)
-
-let make_dep_of_undep env (IndType (indf,realargs)) pj =
- let n = List.length realargs in
- let rec decomp n p =
- if n=0 then p else
- match kind_of_term p with
- | Lambda (_,_,c) -> decomp (n-1) c
- | _ -> decomp (n-1) (applist (lift 1 p, [mkRel 1]))
- in
- let sign,s = decompose_prod_n n pj.uj_type in
- let ind = build_dependent_inductive env indf in
- let s' = mkProd (Anonymous, ind, s) in
- let ccl = lift 1 (decomp n pj.uj_val) in
- let ccl' = mkLambda (Anonymous, ind, ccl) in
- {uj_val=lam_it ccl' sign; uj_type=prod_it s' sign}
-
-(*************************************************************************)
-(* Main pretyping function *)
-
-let pretype_ref isevars env ref =
- let c = constr_of_reference ref in
- make_judge c (Retyping.get_type_of env Evd.empty c)
-
-let pretype_sort = function
- | RProp c -> judge_of_prop_contents c
- | RType _ -> judge_of_new_Type ()
-
-(* [pretype tycon env isevars lvar lmeta cstr] attempts to type [cstr] *)
-(* in environment [env], with existential variables [(evars_of isevars)] and *)
-(* the type constraint tycon *)
-let rec pretype tycon env isevars lvar = function
-
- | RRef (loc,ref) ->
- inh_conv_coerce_to_tycon loc env isevars
- (pretype_ref isevars env ref)
- tycon
-
- | RVar (loc, id) ->
- inh_conv_coerce_to_tycon loc env isevars
- (pretype_id loc env lvar id)
- tycon
-
- | REvar (loc, ev, instopt) ->
- (* Ne faudrait-il pas s'assurer que hyps est bien un
- sous-contexte du contexte courant, et qu'il n'y a pas de Rel "caché" *)
- let hyps = (Evd.map (evars_of isevars) ev).evar_hyps in
- let args = match instopt with
- | None -> instance_from_named_context hyps
- | Some inst -> failwith "Evar subtitutions not implemented" in
- let c = mkEvar (ev, args) in
- let j = (Retyping.get_judgment_of env (evars_of isevars) c) in
- inh_conv_coerce_to_tycon loc env isevars j tycon
-
- | RPatVar (loc,(someta,n)) ->
- anomaly "Found a pattern variable in a rawterm to type"
-
- | RHole (loc,k) ->
- (match tycon with
- | Some ty ->
- { uj_val = new_isevar isevars env (loc,k) ty; uj_type = ty }
- | None -> error_unsolvable_implicit loc env (evars_of isevars) k)
-
- | RRec (loc,fixkind,names,bl,lar,vdef) ->
- let rec type_bl env ctxt = function
- [] -> ctxt
- | (na,None,ty)::bl ->
- let ty' = pretype_type empty_valcon env isevars lvar ty in
- let dcl = (na,None,ty'.utj_val) in
- type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl
- | (na,Some bd,ty)::bl ->
- let ty' = pretype_type empty_valcon env isevars lvar ty in
- let bd' = pretype (mk_tycon ty'.utj_val) env isevars lvar ty in
- let dcl = (na,Some bd'.uj_val,ty'.utj_val) in
- type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl in
- let ctxtv = Array.map (type_bl env empty_rel_context) bl in
- let larj =
- array_map2
- (fun e ar ->
- pretype_type empty_valcon (push_rel_context e env) isevars lvar ar)
- ctxtv lar in
- let lara = Array.map (fun a -> a.utj_val) larj in
- let ftys = array_map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in
- let nbfix = Array.length lar in
- let names = Array.map (fun id -> Name id) names in
- (* Note: bodies are not used by push_rec_types, so [||] is safe *)
- let newenv = push_rec_types (names,ftys,[||]) env in
- let vdefj =
- array_map2_i
- (fun i ctxt def ->
- (* we lift nbfix times the type in tycon, because of
- * the nbfix variables pushed to newenv *)
- let (ctxt,ty) =
- decompose_prod_n_assum (rel_context_length ctxt)
- (lift nbfix ftys.(i)) in
- let nenv = push_rel_context ctxt newenv in
- let j = pretype (mk_tycon ty) nenv isevars lvar def in
- { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt;
- uj_type = it_mkProd_or_LetIn j.uj_type ctxt })
- ctxtv vdef in
- evar_type_fixpoint loc env isevars names ftys vdefj;
- let fixj =
- match fixkind with
- | RFix (vn,i as vni) ->
- let fix = (vni,(names,ftys,Array.map j_val vdefj)) in
- (try check_fix env fix with e -> Stdpp.raise_with_loc loc e);
- make_judge (mkFix fix) ftys.(i)
- | RCoFix i ->
- let cofix = (i,(names,ftys,Array.map j_val vdefj)) in
- (try check_cofix env cofix with e -> Stdpp.raise_with_loc loc e);
- make_judge (mkCoFix cofix) ftys.(i) in
- inh_conv_coerce_to_tycon loc env isevars fixj tycon
-
- | RSort (loc,s) ->
- inh_conv_coerce_to_tycon loc env isevars (pretype_sort s) tycon
-
- | RApp (loc,f,args) ->
- let fj = pretype empty_tycon env isevars lvar f in
- let floc = loc_of_rawconstr f in
- let rec apply_rec env n resj = function
- | [] -> resj
- | c::rest ->
- let argloc = loc_of_rawconstr c in
- let resj = inh_app_fun env isevars resj in
- let resty =
- whd_betadeltaiota env (evars_of isevars) resj.uj_type in
- match kind_of_term resty with
- | Prod (na,c1,c2) ->
- let hj = pretype (mk_tycon c1) env isevars lvar c in
- let newresj =
- { uj_val = applist (j_val resj, [j_val hj]);
- uj_type = subst1 hj.uj_val c2 } in
- apply_rec env (n+1) newresj rest
-
- | _ ->
- let hj = pretype empty_tycon env isevars lvar c in
- error_cant_apply_not_functional_loc
- (join_loc floc argloc) env (evars_of isevars)
- resj [hj]
-
- in let resj = apply_rec env 1 fj args in
- (*
- let apply_one_arg (floc,tycon,jl) c =
- let (dom,rng) = split_tycon floc env isevars tycon in
- let cj = pretype dom env isevars lvar c in
- let rng_tycon = option_app (subst1 cj.uj_val) rng in
- let argloc = loc_of_rawconstr c in
- (join_loc floc argloc,rng_tycon,(argloc,cj)::jl) in
- let _,_,jl =
- List.fold_left apply_one_arg (floc,mk_tycon j.uj_type,[]) args in
- let jl = List.rev jl in
- let resj = inh_apply_rel_list loc env isevars jl (floc,j) tycon in
- *)
- inh_conv_coerce_to_tycon loc env isevars resj tycon
-
- | RLambda(loc,name,c1,c2) ->
- let (name',dom,rng) = split_tycon loc env isevars tycon in
- let dom_valcon = valcon_of_tycon dom in
- let j = pretype_type dom_valcon env isevars lvar c1 in
- let var = (name,None,j.utj_val) in
- let j' = pretype rng (push_rel var env) isevars lvar c2 in
- judge_of_abstraction env name j j'
-
- | RProd(loc,name,c1,c2) ->
- let j = pretype_type empty_valcon env isevars lvar c1 in
- let var = (name,j.utj_val) in
- let env' = push_rel_assum var env in
- let j' = pretype_type empty_valcon env' isevars lvar c2 in
- let resj =
- try judge_of_product env name j j'
- with TypeError _ as e -> Stdpp.raise_with_loc loc e in
- inh_conv_coerce_to_tycon loc env isevars resj tycon
-
- | RLetIn(loc,name,c1,c2) ->
- let j = pretype empty_tycon env isevars lvar c1 in
- let t = Evarutil.refresh_universes j.uj_type in
- let var = (name,Some j.uj_val,t) in
- let tycon = option_app (lift 1) tycon in
- let j' = pretype tycon (push_rel var env) isevars lvar c2 in
- { uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
- uj_type = type_app (subst1 j.uj_val) j'.uj_type }
-
- | RLetTuple (loc,nal,(na,po),c,d) ->
- let cj = pretype empty_tycon env isevars lvar c in
- let (IndType (indf,realargs) as indt) =
- try find_rectype env (evars_of isevars) cj.uj_type
- with Not_found ->
- let cloc = loc_of_rawconstr c in
- error_case_not_inductive_loc cloc env (evars_of isevars) cj
- in
- let cstrs = get_constructors env indf in
- if Array.length cstrs <> 1 then
- user_err_loc (loc,"",str "Destructing let is only for inductive types with one constructor");
- let cs = cstrs.(0) in
- if List.length nal <> cs.cs_nargs then
- user_err_loc (loc,"", str "Destructing let on this type expects " ++ int cs.cs_nargs ++ str " variables");
- let fsign = List.map2 (fun na (_,c,t) -> (na,c,t))
- (List.rev nal) cs.cs_args in
- let env_f = push_rels fsign env in
- (* Make dependencies from arity signature impossible *)
- let arsgn,_ = get_arity env indf in
- let arsgn = List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn in
- let psign = (na,None,build_dependent_inductive env indf)::arsgn in
- let nar = List.length arsgn in
- (match po with
- | Some p ->
- let env_p = push_rels psign env in
- let pj = pretype_type empty_valcon env_p isevars lvar p in
- let ccl = nf_evar (evars_of isevars) pj.utj_val in
- let psign = make_arity_signature env true indf in (* with names *)
- let p = it_mkLambda_or_LetIn ccl psign in
- let inst =
- (Array.to_list cs.cs_concl_realargs)
- @[build_dependent_constructor cs] in
- let lp = lift cs.cs_nargs p in
- let fty = hnf_lam_applist env (evars_of isevars) lp inst in
- let fj = pretype (mk_tycon fty) env_f isevars lvar d in
- let f = it_mkLambda_or_LetIn fj.uj_val fsign in
- let v =
- let mis,_ = dest_ind_family indf in
- let ci = make_default_case_info env LetStyle mis in
- mkCase (ci, p, cj.uj_val,[|f|]) in
- let cs = build_dependent_constructor cs in
- { uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl }
-
- | None ->
- let tycon = option_app (lift cs.cs_nargs) tycon in
- let fj = pretype tycon env_f isevars lvar d in
- let f = it_mkLambda_or_LetIn fj.uj_val fsign in
- let ccl = nf_evar (evars_of isevars) fj.uj_type in
- let ccl =
- if noccur_between 1 cs.cs_nargs ccl then
- lift (- cs.cs_nargs) ccl
- else
- error_cant_find_case_type_loc loc env (evars_of isevars)
- cj.uj_val in
- let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in
- let v =
- let mis,_ = dest_ind_family indf in
- let ci = make_default_case_info env LetStyle mis in
- mkCase (ci, p, cj.uj_val,[|f|] )
- in
- { uj_val = v; uj_type = ccl })
-
- (* Special Case for let constructions to avoid exponential behavior *)
- | ROrderedCase (loc,st,po,c,[|f|],xx) when st <> MatchStyle ->
- let cj = pretype empty_tycon env isevars lvar c in
- let (IndType (indf,realargs) as indt) =
- try find_rectype env (evars_of isevars) cj.uj_type
- with Not_found ->
- let cloc = loc_of_rawconstr c in
- error_case_not_inductive_loc cloc env (evars_of isevars) cj
- in
- let j = match po with
- | Some p ->
- let pj = pretype empty_tycon env isevars lvar p in
- let dep = is_dependent_elimination env pj.uj_type indf in
- let ar =
- arity_of_case_predicate env indf dep (Type (new_univ())) in
- let _ = the_conv_x_leq env isevars pj.uj_type ar in
- let pj = j_nf_evar (evars_of isevars) pj in
- let pj = if dep then pj else make_dep_of_undep env indt pj in
- let (bty,rsty) =
- Indrec.type_rec_branches
- false env (evars_of isevars) indt pj.uj_val cj.uj_val
- in
- if Array.length bty <> 1 then
- error_number_branches_loc
- loc env (evars_of isevars) cj (Array.length bty);
- let fj =
- let tyc = bty.(0) in
- pretype (mk_tycon tyc) env isevars lvar f
- in
- let fv = j_val fj in
- let ft = fj.uj_type in
- check_branches_message loc env isevars cj.uj_val (bty,[|ft|]);
- let v =
- let mis,_ = dest_ind_family indf in
- let ci = make_default_case_info env st mis in
- mkCase (ci, (nf_betaiota pj.uj_val), cj.uj_val,[|fv|])
- in
- { uj_val = v; uj_type = rsty }
-
- | None ->
- (* get type information from type of branches *)
- let expbr = Cases.branch_scheme env isevars false indf in
- if Array.length expbr <> 1 then
- error_number_branches_loc loc env (evars_of isevars)
- cj (Array.length expbr);
- let expti = expbr.(0) in
- let fj = pretype (mk_tycon expti) env isevars lvar f in
- let use_constraint () =
- (* get type information from constraint *)
- (* warning: if the constraint comes from an evar type, it *)
- (* may be Type while Prop or Set would be expected *)
- match tycon with
- | Some pred ->
- let arsgn = make_arity_signature env true indf in
- let pred = lift (List.length arsgn) pred in
- let pred =
- it_mkLambda_or_LetIn (nf_evar (evars_of isevars) pred)
- arsgn in
- false, pred
- | None ->
- let sigma = evars_of isevars in
- error_cant_find_case_type_loc loc env sigma cj.uj_val
- in
- let ok, p =
- try
- let pred =
- Cases.pred_case_ml
- env (evars_of isevars) false indt (0,fj.uj_type)
- in
- if has_undefined_isevars isevars pred then
- use_constraint ()
- else
- true, pred
- with Cases.NotInferable _ ->
- use_constraint ()
- in
- let p = nf_evar (evars_of isevars) p in
- let (bty,rsty) =
- Indrec.type_rec_branches
- false env (evars_of isevars) indt p cj.uj_val
- in
- let _ = option_app (the_conv_x_leq env isevars rsty) tycon in
- let fj =
- if ok then fj
- else pretype (mk_tycon bty.(0)) env isevars lvar f
- in
- let fv = fj.uj_val in
- let ft = fj.uj_type in
- let v =
- let mis,_ = dest_ind_family indf in
- let ci = make_default_case_info env st mis in
- mkCase (ci, (nf_betaiota p), cj.uj_val,[|fv|] )
- in
- { uj_val = v; uj_type = rsty } in
-
- (* Build the LetTuple form for v8 *)
- let c =
- let (ind,params) = dest_ind_family indf in
- let rtntypopt, indnalopt = match po with
- | None -> None, (Anonymous,None)
- | Some p ->
- let pj = pretype empty_tycon env isevars lvar p in
- let dep = is_dependent_elimination env pj.uj_type indf in
- let rec decomp_lam_force n avoid l p =
- (* avoid is not exhaustive ! *)
- if n = 0 then (List.rev l,p,avoid) else
- match p with
- | RLambda (_,(Name id as na),_,c) ->
- decomp_lam_force (n-1) (id::avoid) (na::l) c
- | RLambda (_,(Anonymous as na),_,c) ->
- decomp_lam_force (n-1) avoid (na::l) c
- | _ ->
- let x = Nameops.next_ident_away (id_of_string "x") avoid in
- decomp_lam_force (n-1) (x::avoid) (Name x :: l)
- (* eta-expansion *)
- (RApp (dummy_loc,p, [RVar (dummy_loc,x)])) in
- let (nal,p,avoid) =
- decomp_lam_force (List.length realargs) [] [] p in
- let na,rtntyp,_ =
- if dep then decomp_lam_force 1 avoid [] p
- else [Anonymous],p,[] in
- let intyp =
- if List.for_all
- (function
- | Anonymous -> true
- | Name id -> not (occur_rawconstr id rtntyp)) nal
- then (* No dependency in realargs *)
- None
- else
- let args = List.map (fun _ -> Anonymous) params @ nal in
- Some (dummy_loc,ind,args) in
- (Some rtntyp,(List.hd na,intyp)) in
- let cs = (get_constructors env indf).(0) in
- match indnalopt with
- | (na,None) -> (* Represented as a let *)
- let rec decomp_lam_force n avoid l p =
- if n = 0 then (List.rev l,p) else
- match p with
- | RLambda (_,(Name id as na),_,c) ->
- decomp_lam_force (n-1) (id::avoid) (na::l) c
- | RLambda (_,(Anonymous as na),_,c) ->
- decomp_lam_force (n-1) avoid (na::l) c
- | _ ->
- let x = Nameops.next_ident_away (id_of_string "x") avoid in
- decomp_lam_force (n-1) (x::avoid) (Name x :: l)
- (* eta-expansion *)
- (let a = RVar (dummy_loc,x) in
- match p with
- | RApp (loc,p,l) -> RApp (loc,p,l@[a])
- | _ -> (RApp (dummy_loc,p,[a]))) in
- let (nal,d) = decomp_lam_force cs.cs_nargs [] [] f in
- RLetTuple (loc,nal,(na,rtntypopt),c,d)
- | _ -> (* Represented as a match *)
- let detype_eqn constr construct_nargs branch =
- let name_cons = function
- | Anonymous -> fun l -> l
- | Name id -> fun l -> id::l in
- let make_pat na avoid b ids =
- PatVar (dummy_loc,na),
- name_cons na avoid,name_cons na ids
- in
- let rec buildrec ids patlist avoid n b =
- if n=0 then
- (dummy_loc, ids,
- [PatCstr(dummy_loc, constr, List.rev patlist,Anonymous)],
- b)
- else
- match b with
- | RLambda (_,x,_,b) ->
- let pat,new_avoid,new_ids = make_pat x avoid b ids in
- buildrec new_ids (pat::patlist) new_avoid (n-1) b
-
- | RLetIn (_,x,_,b) ->
- let pat,new_avoid,new_ids = make_pat x avoid b ids in
- buildrec new_ids (pat::patlist) new_avoid (n-1) b
-
- | RCast (_,c,_) -> (* Oui, il y a parfois des cast *)
- buildrec ids patlist avoid n c
-
- | _ -> (* eta-expansion *)
- (* nommage de la nouvelle variable *)
- let id = Nameops.next_ident_away (id_of_string "x") avoid in
- let new_b = RApp (dummy_loc, b, [RVar(dummy_loc,id)])in
- let pat,new_avoid,new_ids =
- make_pat (Name id) avoid new_b ids in
- buildrec new_ids (pat::patlist) new_avoid (n-1) new_b
-
- in
- buildrec [] [] [] construct_nargs branch in
- let eqn = detype_eqn (ind,1) cs.cs_nargs f in
- RCases (loc,(po,ref rtntypopt),[c,ref indnalopt],[eqn])
- in
- xx := Some c;
- (* End building the v8 syntax *)
- j
+(** Miscellaneous interpretation functions *)
+
+let interp_sort = function
+ | RProp c -> Prop c
+ | RType _ -> new_Type_sort ()
+
+let interp_elimination_sort = function
+ | RProp Null -> InProp
+ | RProp Pos -> InSet
+ | RType _ -> InType
- | RIf (loc,c,(na,po),b1,b2) ->
- let cj = pretype empty_tycon env isevars lvar c in
- let (IndType (indf,realargs) as indt) =
- try find_rectype env (evars_of isevars) cj.uj_type
+module type S =
+sig
+
+ module Cases : Cases.S
+
+ (* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
+ val allow_anonymous_refs : bool ref
+
+ (* Generic call to the interpreter from rawconstr to open_constr, leaving
+ unresolved holes as evars and returning the typing contexts of
+ these evars. Work as [understand_gen] for the rest. *)
+
+ val understand_tcc :
+ evar_map -> env -> ?expected_type:types -> rawconstr -> open_constr
+
+ val understand_tcc_evars :
+ evar_defs ref -> env -> typing_constraint -> rawconstr -> constr
+
+ (* More general entry point with evars from ltac *)
+
+ (* Generic call to the interpreter from rawconstr to constr, failing
+ unresolved holes in the rawterm cannot be instantiated.
+
+ In [understand_ltac sigma env ltac_env constraint c],
+
+ sigma : initial set of existential variables (typically dependent subgoals)
+ ltac_env : partial substitution of variables (used for the tactic language)
+ constraint : tell if interpreted as a possibly constrained term or a type
+ *)
+
+ val understand_ltac :
+ evar_map -> env -> var_map * unbound_ltac_var_map ->
+ typing_constraint -> rawconstr -> evar_defs * constr
+
+ (* Standard call to get a constr from a rawconstr, resolving implicit args *)
+
+ val understand : evar_map -> env -> ?expected_type:Term.types ->
+ rawconstr -> constr
+
+ (* Idem but the rawconstr is intended to be a type *)
+
+ val understand_type : evar_map -> env -> rawconstr -> constr
+
+ (* A generalization of the two previous case *)
+
+ val understand_gen : typing_constraint -> evar_map -> env ->
+ rawconstr -> constr
+
+ (* Idem but returns the judgment of the understood term *)
+
+ val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
+
+ (* Idem but do not fail on unresolved evars *)
+
+ val understand_judgment_tcc : evar_defs ref -> env -> rawconstr -> unsafe_judgment
+
+ (*i*)
+ (* Internal of Pretyping...
+ * Unused outside, but useful for debugging
+ *)
+ val pretype :
+ type_constraint -> env -> evar_defs ref ->
+ var_map * (identifier * identifier option) list ->
+ rawconstr -> unsafe_judgment
+
+ val pretype_type :
+ val_constraint -> env -> evar_defs ref ->
+ var_map * (identifier * identifier option) list ->
+ rawconstr -> unsafe_type_judgment
+
+ val pretype_gen :
+ evar_defs ref -> env ->
+ var_map * (identifier * identifier option) list ->
+ typing_constraint -> rawconstr -> constr
+
+ (*i*)
+end
+
+module Pretyping_F (Coercion : Coercion.S) = struct
+
+ module Cases = Cases.Cases_F(Coercion)
+
+ (* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
+ let allow_anonymous_refs = ref false
+
+ let evd_comb0 f isevars =
+ let (evd',x) = f !isevars in
+ isevars := evd';
+ x
+
+ let evd_comb1 f isevars x =
+ let (evd',y) = f !isevars x in
+ isevars := evd';
+ y
+
+ let evd_comb2 f isevars x y =
+ let (evd',z) = f !isevars x y in
+ isevars := evd';
+ z
+
+ let evd_comb3 f isevars x y z =
+ let (evd',t) = f !isevars x y z in
+ isevars := evd';
+ t
+
+ let mt_evd = Evd.empty
+
+ let vect_lift_type = Array.mapi (fun i t -> type_app (lift i) t)
+
+ (* Utilisé pour inférer le prédicat des Cases *)
+ (* Semble exagérement fort *)
+ (* Faudra préférer une unification entre les types de toutes les clauses *)
+ (* et autoriser des ? à rester dans le résultat de l'unification *)
+
+ let evar_type_fixpoint loc env isevars lna lar vdefj =
+ let lt = Array.length vdefj in
+ if Array.length lar = lt then
+ for i = 0 to lt-1 do
+ if not (e_cumul env isevars (vdefj.(i)).uj_type
+ (lift lt lar.(i))) then
+ error_ill_typed_rec_body_loc loc env (evars_of !isevars)
+ i lna vdefj lar
+ done
+
+ let check_branches_message loc env isevars c (explft,lft) =
+ for i = 0 to Array.length explft - 1 do
+ if not (e_cumul env isevars lft.(i) explft.(i)) then
+ let sigma = evars_of !isevars in
+ error_ill_formed_branch_loc loc env sigma c i lft.(i) explft.(i)
+ done
+
+ (* coerce to tycon if any *)
+ let inh_conv_coerce_to_tycon loc env isevars j = function
+ | None -> j
+ | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to loc env) isevars j t
+
+ let push_rels vars env = List.fold_right push_rel vars env
+
+ (*
+ let evar_type_case isevars env ct pt lft p c =
+ let (mind,bty,rslty) = type_case_branches env (evars_of isevars) ct pt p c
+ in check_branches_message isevars env (c,ct) (bty,lft); (mind,rslty)
+ *)
+
+ let strip_meta id = (* For Grammar v7 compatibility *)
+ let s = string_of_id id in
+ if s.[0]='$' then id_of_string (String.sub s 1 (String.length s - 1))
+ else id
+
+ let pretype_id loc env (lvar,unbndltacvars) id =
+ let id = strip_meta id in (* May happen in tactics defined by Grammar *)
+ try
+ let (n,typ) = lookup_rel_id id (rel_context env) in
+ { uj_val = mkRel n; uj_type = type_app (lift n) typ }
+ with Not_found ->
+ try
+ List.assoc id lvar
with Not_found ->
- let cloc = loc_of_rawconstr c in
- error_case_not_inductive_loc cloc env (evars_of isevars) cj in
- let cstrs = get_constructors env indf in
- if Array.length cstrs <> 2 then
- user_err_loc (loc,"",
- str "If is only for inductive types with two constructors");
-
- (* Make dependencies from arity signature impossible *)
- let arsgn,_ = get_arity env indf in
- let arsgn = List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn in
- let nar = List.length arsgn in
- let psign = (na,None,build_dependent_inductive env indf)::arsgn in
- let pred,p = match po with
- | Some p ->
- let env_p = push_rels psign env in
- let pj = pretype_type empty_valcon env_p isevars lvar p in
- let ccl = nf_evar (evars_of isevars) pj.utj_val in
- let pred = it_mkLambda_or_LetIn ccl psign in
- pred, lift (- nar) (beta_applist (pred,[cj.uj_val]))
- | None ->
- let p = match tycon with
- | Some ty -> ty
- | None -> new_isevar isevars env (loc,InternalHole) (new_Type ())
+ try
+ let (_,_,typ) = lookup_named id env in
+ { uj_val = mkVar id; uj_type = typ }
+ with Not_found ->
+ try (* To build a nicer ltac error message *)
+ match List.assoc id unbndltacvars with
+ | None -> user_err_loc (loc,"",
+ str "variable " ++ pr_id id ++ str " should be bound to a term")
+ | Some id0 -> Pretype_errors.error_var_not_found_loc loc id0
+ with Not_found ->
+ error_var_not_found_loc loc id
+
+ (* make a dependent predicate from an undependent one *)
+
+ let make_dep_of_undep env (IndType (indf,realargs)) pj =
+ let n = List.length realargs in
+ let rec decomp n p =
+ if n=0 then p else
+ match kind_of_term p with
+ | Lambda (_,_,c) -> decomp (n-1) c
+ | _ -> decomp (n-1) (applist (lift 1 p, [mkRel 1]))
+ in
+ let sign,s = decompose_prod_n n pj.uj_type in
+ let ind = build_dependent_inductive env indf in
+ let s' = mkProd (Anonymous, ind, s) in
+ let ccl = lift 1 (decomp n pj.uj_val) in
+ let ccl' = mkLambda (Anonymous, ind, ccl) in
+ {uj_val=lam_it ccl' sign; uj_type=prod_it s' sign}
+
+ (*************************************************************************)
+ (* Main pretyping function *)
+
+ let pretype_ref isevars env ref =
+ let c = constr_of_global ref in
+ make_judge c (Retyping.get_type_of env Evd.empty c)
+
+ let pretype_sort = function
+ | RProp c -> judge_of_prop_contents c
+ | RType _ -> judge_of_new_Type ()
+
+ (* [pretype tycon env isevars lvar lmeta cstr] attempts to type [cstr] *)
+ (* in environment [env], with existential variables [(evars_of isevars)] and *)
+ (* the type constraint tycon *)
+ let rec pretype (tycon : type_constraint) env isevars lvar = function
+ | RRef (loc,ref) ->
+ inh_conv_coerce_to_tycon loc env isevars
+ (pretype_ref isevars env ref)
+ tycon
+
+ | RVar (loc, id) ->
+ inh_conv_coerce_to_tycon loc env isevars
+ (pretype_id loc env lvar id)
+ tycon
+
+ | REvar (loc, ev, instopt) ->
+ (* Ne faudrait-il pas s'assurer que hyps est bien un
+ sous-contexte du contexte courant, et qu'il n'y a pas de Rel "caché" *)
+ let hyps = evar_context (Evd.map (evars_of !isevars) ev) in
+ let args = match instopt with
+ | None -> instance_from_named_context hyps
+ | Some inst -> failwith "Evar subtitutions not implemented" in
+ let c = mkEvar (ev, args) in
+ let j = (Retyping.get_judgment_of env (evars_of !isevars) c) in
+ inh_conv_coerce_to_tycon loc env isevars j tycon
+
+ | RPatVar (loc,(someta,n)) ->
+ anomaly "Found a pattern variable in a rawterm to type"
+
+ | RHole (loc,k) ->
+ let ty =
+ match tycon with
+ | Some (None, ty) -> ty
+ | None | Some _ ->
+ e_new_evar isevars env ~src:(loc,InternalHole) (new_Type ()) in
+ { uj_val = e_new_evar isevars env ~src:(loc,k) ty; uj_type = ty }
+
+ | RRec (loc,fixkind,names,bl,lar,vdef) ->
+ let rec type_bl env ctxt = function
+ [] -> ctxt
+ | (na,None,ty)::bl ->
+ let ty' = pretype_type empty_valcon env isevars lvar ty in
+ let dcl = (na,None,ty'.utj_val) in
+ type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl
+ | (na,Some bd,ty)::bl ->
+ let ty' = pretype_type empty_valcon env isevars lvar ty in
+ let bd' = pretype (mk_tycon ty'.utj_val) env isevars lvar ty in
+ let dcl = (na,Some bd'.uj_val,ty'.utj_val) in
+ type_bl (push_rel dcl env) (add_rel_decl dcl ctxt) bl in
+ let ctxtv = Array.map (type_bl env empty_rel_context) bl in
+ let larj =
+ array_map2
+ (fun e ar ->
+ pretype_type empty_valcon (push_rel_context e env) isevars lvar ar)
+ ctxtv lar in
+ let lara = Array.map (fun a -> a.utj_val) larj in
+ let ftys = array_map2 (fun e a -> it_mkProd_or_LetIn a e) ctxtv lara in
+ let nbfix = Array.length lar in
+ let names = Array.map (fun id -> Name id) names in
+ (* Note: bodies are not used by push_rec_types, so [||] is safe *)
+ let newenv = push_rec_types (names,ftys,[||]) env in
+ let vdefj =
+ array_map2_i
+ (fun i ctxt def ->
+ (* we lift nbfix times the type in tycon, because of
+ * the nbfix variables pushed to newenv *)
+ let (ctxt,ty) =
+ decompose_prod_n_assum (rel_context_length ctxt)
+ (lift nbfix ftys.(i)) in
+ let nenv = push_rel_context ctxt newenv in
+ let j = pretype (mk_tycon ty) nenv isevars lvar def in
+ { uj_val = it_mkLambda_or_LetIn j.uj_val ctxt;
+ uj_type = it_mkProd_or_LetIn j.uj_type ctxt })
+ ctxtv vdef in
+ evar_type_fixpoint loc env isevars names ftys vdefj;
+ let fixj =
+ match fixkind with
+ | RFix (vn,i) ->
+ let fix = ((Array.map fst vn, i),(names,ftys,Array.map j_val vdefj)) in
+ (try check_fix env fix with e -> Stdpp.raise_with_loc loc e);
+ make_judge (mkFix fix) ftys.(i)
+ | RCoFix i ->
+ let cofix = (i,(names,ftys,Array.map j_val vdefj)) in
+ (try check_cofix env cofix with e -> Stdpp.raise_with_loc loc e);
+ make_judge (mkCoFix cofix) ftys.(i) in
+ inh_conv_coerce_to_tycon loc env isevars fixj tycon
+
+ | RSort (loc,s) ->
+ inh_conv_coerce_to_tycon loc env isevars (pretype_sort s) tycon
+
+ | RApp (loc,f,args) ->
+ let length = List.length args in
+ let ftycon =
+ match tycon with
+ None -> None
+ | Some (None, ty) -> mk_abstr_tycon length ty
+ | Some (Some (init, cur), ty) ->
+ Some (Some (length + init, length + cur), ty)
+ in
+ let fj = pretype empty_tycon env isevars lvar f in
+ let floc = loc_of_rawconstr f in
+ let rec apply_rec env n resj tycon = function
+ | [] -> resj
+ | c::rest ->
+ let argloc = loc_of_rawconstr c in
+ let resj = evd_comb1 (Coercion.inh_app_fun env) isevars resj in
+ let resty = whd_betadeltaiota env (evars_of !isevars) resj.uj_type in
+ match kind_of_term resty with
+ | Prod (na,c1,c2) ->
+ let hj = pretype (mk_tycon c1) env isevars lvar c in
+ let value, typ = applist (j_val resj, [j_val hj]), subst1 hj.uj_val c2 in
+ let typ' = nf_isevar !isevars typ in
+ let tycon =
+ option_app
+ (fun (abs, ty) ->
+ match abs with
+ None ->
+ isevars := Coercion.inh_conv_coerces_to loc env !isevars typ'
+ (abs, ty);
+ (abs, ty)
+ | Some (init, cur) ->
+ isevars := Coercion.inh_conv_coerces_to loc env !isevars typ'
+ (abs, ty);
+ (Some (init, pred cur), ty))
+ tycon
+ in
+ apply_rec env (n+1)
+ { uj_val = nf_isevar !isevars value;
+ uj_type = nf_isevar !isevars typ' }
+ (option_app (fun (abs, c) -> abs, nf_isevar !isevars c) tycon) rest
+
+ | _ ->
+ let hj = pretype empty_tycon env isevars lvar c in
+ error_cant_apply_not_functional_loc
+ (join_loc floc argloc) env (evars_of !isevars)
+ resj [hj]
+ in
+ let ftycon = option_app (lift_abstr_tycon_type (-1)) ftycon in
+ let resj = j_nf_evar (evars_of !isevars) (apply_rec env 1 fj ftycon args) in
+ let resj =
+ match kind_of_term resj.uj_val with
+ | App (f,args) when isInd f ->
+ let sigma = evars_of !isevars in
+ let t = Retyping.type_of_applied_inductive env sigma (destInd f) args in
+ let s = snd (splay_arity env sigma t) in
+ on_judgment_type (set_inductive_level env s) resj
+ (* Rem: no need to send sigma: no head evar, it's an arity *)
+ | _ -> resj in
+ inh_conv_coerce_to_tycon loc env isevars resj tycon
+
+ | RLambda(loc,name,c1,c2) ->
+ let (name',dom,rng) = evd_comb1 (split_tycon loc env) isevars tycon in
+ let dom_valcon = valcon_of_tycon dom in
+ let j = pretype_type dom_valcon env isevars lvar c1 in
+ let var = (name,None,j.utj_val) in
+ let j' = pretype rng (push_rel var env) isevars lvar c2 in
+ judge_of_abstraction env name j j'
+
+ | RProd(loc,name,c1,c2) ->
+ let j = pretype_type empty_valcon env isevars lvar c1 in
+ let var = (name,j.utj_val) in
+ let env' = push_rel_assum var env in
+ let j' = pretype_type empty_valcon env' isevars lvar c2 in
+ let resj =
+ try judge_of_product env name j j'
+ with TypeError _ as e -> Stdpp.raise_with_loc loc e in
+ inh_conv_coerce_to_tycon loc env isevars resj tycon
+
+ | RLetIn(loc,name,c1,c2) ->
+ let j = pretype empty_tycon env isevars lvar c1 in
+ let t = refresh_universes j.uj_type in
+ let var = (name,Some j.uj_val,t) in
+ let tycon = lift_tycon 1 tycon in
+ let j' = pretype tycon (push_rel var env) isevars lvar c2 in
+ { uj_val = mkLetIn (name, j.uj_val, t, j'.uj_val) ;
+ uj_type = subst1 j.uj_val j'.uj_type }
+
+ | RLetTuple (loc,nal,(na,po),c,d) ->
+ let cj = pretype empty_tycon env isevars lvar c in
+ let (IndType (indf,realargs)) =
+ try find_rectype env (evars_of !isevars) cj.uj_type
+ with Not_found ->
+ let cloc = loc_of_rawconstr c in
+ error_case_not_inductive_loc cloc env (evars_of !isevars) cj
+ in
+ let cstrs = get_constructors env indf in
+ if Array.length cstrs <> 1 then
+ user_err_loc (loc,"",str "Destructing let is only for inductive types with one constructor");
+ let cs = cstrs.(0) in
+ if List.length nal <> cs.cs_nargs then
+ user_err_loc (loc,"", str "Destructing let on this type expects " ++ int cs.cs_nargs ++ str " variables");
+ let fsign = List.map2 (fun na (_,c,t) -> (na,c,t))
+ (List.rev nal) cs.cs_args in
+ let env_f = push_rels fsign env in
+ (* Make dependencies from arity signature impossible *)
+ let arsgn =
+ let arsgn,_ = get_arity env indf in
+ if not !allow_anonymous_refs then
+ List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn
+ else arsgn
in
- it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in
- let f cs b =
- let n = rel_context_length cs.cs_args in
- let pi = liftn n 2 pred in
- let pi = beta_applist (pi, [build_dependent_constructor cs]) in
- let csgn = List.map (fun (_,b,t) -> (Anonymous,b,t)) cs.cs_args in
- let env_c = push_rels csgn env in
- let bj = pretype (Some pi) env_c isevars lvar b in
- it_mkLambda_or_LetIn bj.uj_val cs.cs_args in
- let b1 = f cstrs.(0) b1 in
- let b2 = f cstrs.(1) b2 in
- let pred = nf_evar (evars_of isevars) pred in
- let p = nf_evar (evars_of isevars) p in
- let v =
- let mis,_ = dest_ind_family indf in
- let ci = make_default_case_info env IfStyle mis in
- mkCase (ci, pred, cj.uj_val, [|b1;b2|])
- in
- { uj_val = v; uj_type = p }
-
- | ROrderedCase (loc,st,po,c,lf,x) ->
- let isrec = (st = MatchStyle) in
- let cj = pretype empty_tycon env isevars lvar c in
- let (IndType (indf,realargs) as indt) =
- try find_rectype env (evars_of isevars) cj.uj_type
- with Not_found ->
- let cloc = loc_of_rawconstr c in
- error_case_not_inductive_loc cloc env (evars_of isevars) cj in
- let (dep,pj) = match po with
- | Some p ->
- let pj = pretype empty_tycon env isevars lvar p in
- let dep = is_dependent_elimination env pj.uj_type indf in
- let ar =
- arity_of_case_predicate env indf dep (Type (new_univ())) in
- let _ = the_conv_x_leq env isevars pj.uj_type ar in
- (dep, pj)
- | None ->
- (* get type information from type of branches *)
- let expbr = Cases.branch_scheme env isevars isrec indf in
- let rec findtype i =
- if i >= Array.length lf
- then
- (* get type information from constraint *)
- (* warning: if the constraint comes from an evar type, it *)
- (* may be Type while Prop or Set would be expected *)
- match tycon with
- | Some pred ->
- let arsgn = make_arity_signature env true indf in
- let pred = lift (List.length arsgn) pred in
- let pred =
- it_mkLambda_or_LetIn (nf_evar (evars_of isevars) pred)
- arsgn in
- (true,
- Retyping.get_judgment_of env (evars_of isevars) pred)
- | None ->
- let sigma = evars_of isevars in
- error_cant_find_case_type_loc loc env sigma cj.uj_val
- else
- try
- let expti = expbr.(i) in
- let fj =
- pretype (mk_tycon expti) env isevars lvar lf.(i) in
- let pred =
- Cases.pred_case_ml (* eta-expanse *)
- env (evars_of isevars) isrec indt (i,fj.uj_type) in
- if has_undefined_isevars isevars pred then findtype (i+1)
- else
- let pty =
- Retyping.get_type_of env (evars_of isevars) pred in
- let pj = { uj_val = pred; uj_type = pty } in
-(*
- let _ = option_app (the_conv_x_leq env isevars pred) tycon
- in
-*)
- (true,pj)
- with Cases.NotInferable _ -> findtype (i+1) in
- findtype 0
- in
- let pj = j_nf_evar (evars_of isevars) pj in
- let pj = if dep then pj else make_dep_of_undep env indt pj in
- let (bty,rsty) =
- Indrec.type_rec_branches
- isrec env (evars_of isevars) indt pj.uj_val cj.uj_val in
- let _ = option_app (the_conv_x_leq env isevars rsty) tycon in
- if Array.length bty <> Array.length lf then
- error_number_branches_loc loc env (evars_of isevars)
- cj (Array.length bty)
- else
- let lfj =
- array_map2
- (fun tyc f -> pretype (mk_tycon tyc) env isevars lvar f) bty
- lf in
- let lfv = Array.map j_val lfj in
- let lft = Array.map (fun j -> j.uj_type) lfj in
- check_branches_message loc env isevars cj.uj_val (bty,lft);
- let v =
- if isrec
- then
- transform_rec loc env (evars_of isevars)(pj,cj.uj_val,lfv) indt
- else
+ let psign = (na,None,build_dependent_inductive env indf)::arsgn in
+ let nar = List.length arsgn in
+ (match po with
+ | Some p ->
+ let env_p = push_rels psign env in
+ let pj = pretype_type empty_valcon env_p isevars lvar p in
+ let ccl = nf_evar (evars_of !isevars) pj.utj_val in
+ let psign = make_arity_signature env true indf in (* with names *)
+ let p = it_mkLambda_or_LetIn ccl psign in
+ let inst =
+ (Array.to_list cs.cs_concl_realargs)
+ @[build_dependent_constructor cs] in
+ let lp = lift cs.cs_nargs p in
+ let fty = hnf_lam_applist env (evars_of !isevars) lp inst in
+ let fj = pretype (mk_tycon fty) env_f isevars lvar d in
+ let f = it_mkLambda_or_LetIn fj.uj_val fsign in
+ let v =
+ let mis,_ = dest_ind_family indf in
+ let ci = make_default_case_info env LetStyle mis in
+ mkCase (ci, p, cj.uj_val,[|f|]) in
+ { uj_val = v; uj_type = substl (realargs@[cj.uj_val]) ccl }
+
+ | None ->
+ let tycon = lift_tycon cs.cs_nargs tycon in
+ let fj = pretype tycon env_f isevars lvar d in
+ let f = it_mkLambda_or_LetIn fj.uj_val fsign in
+ let ccl = nf_evar (evars_of !isevars) fj.uj_type in
+ let ccl =
+ if noccur_between 1 cs.cs_nargs ccl then
+ lift (- cs.cs_nargs) ccl
+ else
+ error_cant_find_case_type_loc loc env (evars_of !isevars)
+ cj.uj_val in
+ let p = it_mkLambda_or_LetIn (lift (nar+1) ccl) psign in
+ let v =
+ let mis,_ = dest_ind_family indf in
+ let ci = make_default_case_info env LetStyle mis in
+ mkCase (ci, p, cj.uj_val,[|f|] )
+ in
+ { uj_val = v; uj_type = ccl })
+
+ | RIf (loc,c,(na,po),b1,b2) ->
+ let cj = pretype empty_tycon env isevars lvar c in
+ let (IndType (indf,realargs)) =
+ try find_rectype env (evars_of !isevars) cj.uj_type
+ with Not_found ->
+ let cloc = loc_of_rawconstr c in
+ error_case_not_inductive_loc cloc env (evars_of !isevars) cj in
+ let cstrs = get_constructors env indf in
+ if Array.length cstrs <> 2 then
+ user_err_loc (loc,"",
+ str "If is only for inductive types with two constructors");
+
+ let arsgn =
+ let arsgn,_ = get_arity env indf in
+ if not !allow_anonymous_refs then
+ (* Make dependencies from arity signature impossible *)
+ List.map (fun (_,b,t) -> (Anonymous,b,t)) arsgn
+ else arsgn
+ in
+ let nar = List.length arsgn in
+ let psign = (na,None,build_dependent_inductive env indf)::arsgn in
+ let pred,p = match po with
+ | Some p ->
+ let env_p = push_rels psign env in
+ let pj = pretype_type empty_valcon env_p isevars lvar p in
+ let ccl = nf_evar (evars_of !isevars) pj.utj_val in
+ let pred = it_mkLambda_or_LetIn ccl psign in
+ let typ = lift (- nar) (beta_applist (pred,[cj.uj_val])) in
+ let jtyp = inh_conv_coerce_to_tycon loc env isevars {uj_val = pred;
+ uj_type = typ} tycon
+ in
+ jtyp.uj_val, jtyp.uj_type
+ | None ->
+ let p = match tycon with
+ | Some (None, ty) -> ty
+ | None | Some _ ->
+ e_new_evar isevars env ~src:(loc,InternalHole) (new_Type ())
+ in
+ it_mkLambda_or_LetIn (lift (nar+1) p) psign, p in
+ let pred = nf_evar (evars_of !isevars) pred in
+ let p = nf_evar (evars_of !isevars) p in
+ (* msgnl (str "Pred is: " ++ Termops.print_constr_env env pred);*)
+ let f cs b =
+ let n = rel_context_length cs.cs_args in
+ let pi = lift n pred in (* liftn n 2 pred ? *)
+ let pi = beta_applist (pi, [build_dependent_constructor cs]) in
+ let csgn =
+ if not !allow_anonymous_refs then
+ List.map (fun (_,b,t) -> (Anonymous,b,t)) cs.cs_args
+ else
+ List.map
+ (fun (n, b, t) ->
+ match n with
+ Name _ -> (n, b, t)
+ | Anonymous -> (Name (id_of_string "H"), b, t))
+ cs.cs_args
+ in
+ let env_c = push_rels csgn env in
+(* msgnl (str "Pi is: " ++ Termops.print_constr_env env_c pi); *)
+ let bj = pretype (mk_tycon pi) env_c isevars lvar b in
+ it_mkLambda_or_LetIn bj.uj_val cs.cs_args in
+ let b1 = f cstrs.(0) b1 in
+ let b2 = f cstrs.(1) b2 in
+ let v =
let mis,_ = dest_ind_family indf in
- let ci = make_default_case_info env st mis in
- mkCase (ci, (nf_betaiota pj.uj_val), cj.uj_val,
- Array.map (fun j-> j.uj_val) lfj)
- in
- (* Build the Cases form for v8 *)
- let c =
- let (ind,params) = dest_ind_family indf in
- let (mib,mip) = lookup_mind_specif env ind in
- let recargs = mip.mind_recargs in
- let mI = mkInd ind in
- let nconstr = Array.length mip.mind_consnames in
- let tyi = snd ind in
- if isrec && mis_is_recursive_subset [tyi] recargs then
- Some (Detyping.detype (false,env)
- (ids_of_context env) (names_of_rel_context env)
- (nf_evar (evars_of isevars) v))
- else
- (* Translate into a "match ... with" *)
- let rtntypopt, indnalopt = match po with
- | None -> None, (Anonymous,None)
- | Some p ->
- let rec decomp_lam_force n avoid l p =
- (* avoid is not exhaustive ! *)
- if n = 0 then (List.rev l,p,avoid) else
- match p with
- | RLambda (_,(Name id as na),_,c) ->
- decomp_lam_force (n-1) (id::avoid) (na::l) c
- | RLambda (_,(Anonymous as na),_,c) ->
- decomp_lam_force (n-1) avoid (na::l) c
- | _ ->
- let x = Nameops.next_ident_away (id_of_string "x") avoid in
- decomp_lam_force (n-1) (x::avoid) (Name x :: l)
- (* eta-expansion *)
- (RApp (dummy_loc,p, [RVar (dummy_loc,x)])) in
- let (nal,p,avoid) =
- decomp_lam_force (List.length realargs) [] [] p in
- let na,rtntyopt,_ =
- if dep then decomp_lam_force 1 avoid [] p
- else [Anonymous],p,[] in
- let intyp =
- if nal=[] then None else
- let args = List.map (fun _ -> Anonymous) params @ nal in
- Some (dummy_loc,ind,args) in
- (Some rtntyopt,(List.hd na,intyp)) in
- let rawbranches =
- array_map3 (fun bj b cstr ->
- let rec strip n r = if n=0 then r else
- match r with
- | RLambda (_,_,_,t) -> strip (n-1) t
- | RLetIn (_,_,_,t) -> strip (n-1) t
- | _ -> assert false in
- let n = rel_context_length cstr.cs_args in
- try
- let _,ccl = decompose_lam_n_assum n bj.uj_val in
- if noccur_between 1 n ccl then Some (strip n b) else None
- with _ -> (* Not eta-expanded or not reduced *) None)
- lfj lf (get_constructors env indf) in
- if st = IfStyle & snd indnalopt = None
- & rawbranches.(0) <> None && rawbranches.(1) <> None then
- (* Translate into a "if ... then ... else" *)
- (* TODO: translate into a "if" even if po is dependent *)
- Some (RIf (loc,c,(fst indnalopt,rtntypopt),
- out_some rawbranches.(0),out_some rawbranches.(1)))
- else
- let detype_eqn constr construct_nargs branch =
- let name_cons = function
- | Anonymous -> fun l -> l
- | Name id -> fun l -> id::l in
- let make_pat na avoid b ids =
- PatVar (dummy_loc,na),
- name_cons na avoid,name_cons na ids
- in
- let rec buildrec ids patlist avoid n b =
- if n=0 then
- (dummy_loc, ids,
- [PatCstr(dummy_loc, constr, List.rev patlist,Anonymous)],
- b)
- else
- match b with
- | RLambda (_,x,_,b) ->
- let pat,new_avoid,new_ids = make_pat x avoid b ids in
- buildrec new_ids (pat::patlist) new_avoid (n-1) b
-
- | RLetIn (_,x,_,b) ->
- let pat,new_avoid,new_ids = make_pat x avoid b ids in
- buildrec new_ids (pat::patlist) new_avoid (n-1) b
-
- | RCast (_,c,_) -> (* Oui, il y a parfois des cast *)
- buildrec ids patlist avoid n c
-
- | _ -> (* eta-expansion *)
- (* nommage de la nouvelle variable *)
- let id = Nameops.next_ident_away (id_of_string "x") avoid in
- let new_b = RApp (dummy_loc, b, [RVar(dummy_loc,id)])in
- let pat,new_avoid,new_ids =
- make_pat (Name id) avoid new_b ids in
- buildrec new_ids (pat::patlist) new_avoid (n-1) new_b
-
- in
- buildrec [] [] [] construct_nargs branch in
- let (mib,mip) = Inductive.lookup_mind_specif (Global.env()) ind in
- let get_consnarg j =
- let typi = mis_nf_constructor_type (ind,mib,mip) (j+1) in
- let _,t = decompose_prod_n_assum mip.mind_nparams typi in
- List.rev (fst (decompose_prod_assum t)) in
- let consnargs = Array.init (Array.length mip.mind_consnames) get_consnarg in
- let consnargsl = Array.map List.length consnargs in
- let constructs = Array.init (Array.length lf) (fun i -> (ind,i+1)) in
- let eqns = array_map3 detype_eqn constructs consnargsl lf in
- Some (RCases (loc,(po,ref rtntypopt),[c,ref indnalopt],Array.to_list eqns)) in
- x := c;
- (* End build the Cases form for v8 *)
- { uj_val = v;
- uj_type = rsty }
-
- | RCases (loc,po,tml,eqns) ->
- Cases.compile_cases loc
- ((fun vtyc env -> pretype vtyc env isevars lvar),isevars)
- tycon env (* loc *) (po,tml,eqns)
-
- | RCast(loc,c,t) ->
- let tj = pretype_type empty_tycon env isevars lvar t in
- let cj = pretype (mk_tycon tj.utj_val) env isevars lvar c in
- (* User Casts are for helping pretyping, experimentally not to be kept*)
- (* ... except for Correctness *)
- let v = mkCast (cj.uj_val, tj.utj_val) in
- let cj = { uj_val = v; uj_type = tj.utj_val } in
- inh_conv_coerce_to_tycon loc env isevars cj tycon
-
- | RDynamic (loc,d) ->
- if (tag d) = "constr" then
- let c = constr_out d in
- let j = (Retyping.get_judgment_of env (evars_of isevars) c) in
+ let ci = make_default_case_info env IfStyle mis in
+ mkCase (ci, pred, cj.uj_val, [|b1;b2|])
+ in
+ { uj_val = v; uj_type = p }
+
+ | RCases (loc,po,tml,eqns) ->
+ Cases.compile_cases loc
+ ((fun vtyc env -> pretype vtyc env isevars lvar),isevars)
+ tycon env (* loc *) (po,tml,eqns)
+
+ | RCast(loc,c,k,t) ->
+ let tj = pretype_type empty_valcon env isevars lvar t in
+ let cj = pretype (mk_tycon tj.utj_val) env isevars lvar c in
+ (* User Casts are for helping pretyping, experimentally not to be kept*)
+ (* ... except for Correctness *)
+ let v = mkCast (cj.uj_val, k, tj.utj_val) in
+ let cj = { uj_val = v; uj_type = tj.utj_val } in
+ inh_conv_coerce_to_tycon loc env isevars cj tycon
+
+ | RDynamic (loc,d) ->
+ if (tag d) = "constr" then
+ let c = constr_out d in
+ let j = (Retyping.get_judgment_of env (evars_of !isevars) c) in
+ j
+ (*inh_conv_coerce_to_tycon loc env isevars j tycon*)
+ else
+ user_err_loc (loc,"pretype",(str "Not a constr tagged Dynamic"))
+
+ (* [pretype_type valcon env isevars lvar c] coerces [c] into a type *)
+ and pretype_type valcon env isevars lvar = function
+ | RHole loc ->
+ (match valcon with
+ | Some v ->
+ let s =
+ let sigma = evars_of !isevars in
+ let t = Retyping.get_type_of env sigma v in
+ match kind_of_term (whd_betadeltaiota env sigma t) with
+ | Sort s -> s
+ | Evar v when is_Type (existential_type sigma v) ->
+ evd_comb1 (define_evar_as_sort) isevars v
+ | _ -> anomaly "Found a type constraint which is not a type"
+ in
+ { utj_val = v;
+ utj_type = s }
+ | None ->
+ let s = new_Type_sort () in
+ { utj_val = e_new_evar isevars env ~src:loc (mkSort s);
+ utj_type = s})
+ | c ->
+ let j = pretype empty_tycon env isevars lvar c in
+ let loc = loc_of_rawconstr c in
+ let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) isevars j in
+ match valcon with
+ | None -> tj
+ | Some v ->
+ if e_cumul env isevars v tj.utj_val then tj
+ else
+ error_unexpected_type_loc
+ (loc_of_rawconstr c) env (evars_of !isevars) tj.utj_val v
+
+ let pretype_gen isevars env lvar kind c =
+ let c' = match kind with
+ | OfType exptyp ->
+ let tycon = match exptyp with None -> empty_tycon | Some t -> mk_tycon t in
+ (pretype tycon env isevars lvar c).uj_val
+ | IsType ->
+ (pretype_type empty_valcon env isevars lvar c).utj_val in
+ nf_evar (evars_of !isevars) c'
+
+ (* [check_evars] fails if some unresolved evar remains *)
+ (* it assumes that the defined existentials have already been substituted
+ (should be done in unsafe_infer and unsafe_infer_type) *)
+
+ let check_evars env initial_sigma isevars c =
+ let sigma = evars_of !isevars in
+ let rec proc_rec c =
+ match kind_of_term c with
+ | Evar (ev,args) ->
+ assert (Evd.in_dom sigma ev);
+ if not (Evd.in_dom initial_sigma ev) then
+ let (loc,k) = evar_source ev !isevars in
+ error_unsolvable_implicit loc env sigma k
+ | _ -> iter_constr proc_rec c
+ in
+ proc_rec c(*;
+ let (_,pbs) = get_conv_pbs !isevars (fun _ -> true) in
+ if pbs <> [] then begin
+ pperrnl
+ (str"TYPING OF "++Termops.print_constr_env env c++fnl()++
+ prlist_with_sep fnl
+ (fun (pb,c1,c2) ->
+ Termops.print_constr c1 ++
+ (if pb=Reduction.CUMUL then str " <="++ spc()
+ else str" =="++spc()) ++
+ Termops.print_constr c2)
+ pbs ++ fnl())
+ end*)
+
+ (* TODO: comment faire remonter l'information si le typage a resolu des
+ variables du sigma original. il faudrait que la fonction de typage
+ retourne aussi le nouveau sigma...
+ *)
+
+ let understand_judgment sigma env c =
+ let isevars = ref (create_evar_defs sigma) in
+ let j = pretype empty_tycon env isevars ([],[]) c in
+ let j = j_nf_evar (evars_of !isevars) j in
+ check_evars env sigma isevars (mkCast(j.uj_val,DEFAULTcast, j.uj_type));
j
- (*inh_conv_coerce_to_tycon loc env isevars j tycon*)
- else
- user_err_loc (loc,"pretype",(str "Not a constr tagged Dynamic"))
-
-(* [pretype_type valcon env isevars lvar c] coerces [c] into a type *)
-and pretype_type valcon env isevars lvar = function
- | RHole loc ->
- (match valcon with
- | Some v ->
- let s =
- let sigma = evars_of isevars in
- let t = Retyping.get_type_of env sigma v in
- match kind_of_term (whd_betadeltaiota env sigma t) with
- | Sort s -> s
- | Evar v when is_Type (existential_type sigma v) ->
- define_evar_as_sort isevars v
- | _ -> anomaly "Found a type constraint which is not a type"
- in
- { utj_val = v;
- utj_type = s }
- | None ->
- let s = new_Type_sort () in
- { utj_val = new_isevar isevars env loc (mkSort s);
- utj_type = s})
- | c ->
- let j = pretype empty_tycon env isevars lvar c in
- let tj = inh_coerce_to_sort env isevars j in
- match valcon with
- | None -> tj
- | Some v ->
- if the_conv_x_leq env isevars v tj.utj_val then tj
- else
- error_unexpected_type_loc
- (loc_of_rawconstr c) env (evars_of isevars) tj.utj_val v
-
-
-let unsafe_infer tycon isevars env lvar constr =
- let j = pretype tycon env isevars lvar constr in
- j_nf_evar (evars_of isevars) j
-
-let unsafe_infer_type valcon isevars env lvar constr =
- let tj = pretype_type valcon env isevars lvar constr in
- tj_nf_evar (evars_of isevars) tj
-
-(* If fail_evar is false, [process_evars] builds a meta_map with the
- unresolved Evar that were not in initial sigma; otherwise it fail
- on the first unresolved Evar not already in the initial sigma. *)
-(* [fail_evar] says how to process unresolved evars:
- * true -> raise an error message
- * false -> convert them into new Metas (casted with their type)
- *)
-(* assumes the defined existentials have been replaced in c (should be
- done in unsafe_infer and unsafe_infer_type) *)
-let check_evars fail_evar env initial_sigma isevars c =
- let sigma = evars_of isevars in
- let rec proc_rec c =
- match kind_of_term c with
- | Evar (ev,args as k) ->
- assert (Evd.in_dom sigma ev);
- if not (Evd.in_dom initial_sigma ev) then
- (if fail_evar then
- let (loc,k) = evar_source ev isevars in
- error_unsolvable_implicit loc env sigma k)
- | _ -> iter_constr proc_rec c
- in
- proc_rec c
-
-(* TODO: comment faire remonter l'information si le typage a resolu des
- variables du sigma original. il faudrait que la fonction de typage
- retourne aussi le nouveau sigma...
-*)
-
-(* constr with holes *)
-type open_constr = evar_map * constr
-
-let ise_resolve_casted_gen fail_evar sigma env lvar typ c =
- let isevars = create_evar_defs sigma in
- let j = unsafe_infer (mk_tycon typ) isevars env lvar c in
- check_evars fail_evar env sigma isevars (mkCast(j.uj_val,j.uj_type));
- (evars_of isevars, j)
-
-let ise_resolve_casted sigma env typ c =
- ise_resolve_casted_gen true sigma env ([],[]) typ c
-
-(* Raw calls to the unsafe inference machine: boolean says if we must fail
- on unresolved evars, or replace them by Metas; the unsafe_judgment list
- allows us to extend env with some bindings *)
-let ise_infer_gen fail_evar sigma env lvar exptyp c =
- let tycon = match exptyp with None -> empty_tycon | Some t -> mk_tycon t in
- let isevars = create_evar_defs sigma in
- let j = unsafe_infer tycon isevars env lvar c in
- check_evars fail_evar env sigma isevars (mkCast(j.uj_val,j.uj_type));
- (evars_of isevars, j)
-
-let ise_infer_type_gen fail_evar sigma env lvar c =
- let isevars = create_evar_defs sigma in
- let tj = unsafe_infer_type empty_valcon isevars env lvar c in
- check_evars fail_evar env sigma isevars tj.utj_val;
- (evars_of isevars, tj)
-type var_map = (identifier * unsafe_judgment) list
+ let understand_judgment_tcc isevars env c =
+ let j = pretype empty_tycon env isevars ([],[]) c in
+ let sigma = evars_of !isevars in
+ let j = j_nf_evar sigma j in
+ j
-let understand_judgment sigma env c =
- snd (ise_infer_gen true sigma env ([],[]) None c)
+ (* Raw calls to the unsafe inference machine: boolean says if we must
+ fail on unresolved evars; the unsafe_judgment list allows us to
+ extend env with some bindings *)
-let understand_type_judgment sigma env c =
- snd (ise_infer_type_gen true sigma env ([],[]) c)
+ let ise_pretype_gen fail_evar sigma env lvar kind c =
+ let isevars = ref (Evd.create_evar_defs sigma) in
+ let c = pretype_gen isevars env lvar kind c in
+ if fail_evar then check_evars env sigma isevars c;
+ !isevars, c
-let understand sigma env c =
- let _, c = ise_infer_gen true sigma env ([],[]) None c in
- c.uj_val
+ (** Entry points of the high-level type synthesis algorithm *)
-let understand_type sigma env c =
- let _,c = ise_infer_type_gen true sigma env ([],[]) c in
- c.utj_val
+ let understand_gen kind sigma env c =
+ snd (ise_pretype_gen true sigma env ([],[]) kind c)
-let understand_gen_ltac sigma env lvar ~expected_type:exptyp c =
- let _, c = ise_infer_gen true sigma env lvar exptyp c in
- c.uj_val
+ let understand sigma env ?expected_type:exptyp c =
+ snd (ise_pretype_gen true sigma env ([],[]) (OfType exptyp) c)
-let understand_gen sigma env lvar ~expected_type:exptyp c =
- let _, c = ise_infer_gen true sigma env (lvar,[]) exptyp c in
- c.uj_val
+ let understand_type sigma env c =
+ snd (ise_pretype_gen true sigma env ([],[]) IsType c)
-let understand_gen_tcc sigma env lvar exptyp c =
- let metamap, c = ise_infer_gen false sigma env (lvar,[]) exptyp c in
- metamap, c.uj_val
+ let understand_ltac sigma env lvar kind c =
+ ise_pretype_gen false sigma env lvar kind c
+
+ let understand_tcc_evars isevars env kind c =
+ pretype_gen isevars env ([],[]) kind c
-let interp_sort = function
- | RProp c -> Prop c
- | RType _ -> new_Type_sort ()
+ let understand_tcc sigma env ?expected_type:exptyp c =
+ let ev, t = ise_pretype_gen false sigma env ([],[]) (OfType exptyp) c in
+ Evd.evars_of ev, t
+end
-let interp_elimination_sort = function
- | RProp Null -> InProp
- | RProp Pos -> InSet
- | RType _ -> InType
+module Default : S = Pretyping_F(Coercion.Default)
diff --git a/pretyping/pretyping.mli b/pretyping/pretyping.mli
index 4357e504..7bb8c374 100644
--- a/pretyping/pretyping.mli
+++ b/pretyping/pretyping.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: pretyping.mli,v 1.28.2.1 2004/07/16 19:30:46 herbelin Exp $ i*)
+(*i $Id: pretyping.mli 8688 2006-04-07 15:08:12Z msozeau $ i*)
(*i*)
open Names
@@ -18,69 +18,97 @@ open Rawterm
open Evarutil
(*i*)
-type var_map = (identifier * unsafe_judgment) list
-
-(* constr with holes *)
-type open_constr = evar_map * constr
-
-
-(* Generic call to the interpreter from rawconstr to constr, failing
- unresolved holes in the rawterm cannot be instantiated.
-
- In [understand_gen sigma env varmap typopt raw],
-
- sigma : initial set of existential variables (typically dependent subgoals)
- varmap : partial subtitution of variables (used for the tactic language)
- metamap : partial subtitution of meta (used for the tactic language)
- typopt : is not None, this is the expected type for raw (used to define evars)
-*)
-val understand_gen :
- evar_map -> env -> var_map
- -> expected_type:(constr option) -> rawconstr -> constr
+type typing_constraint = OfType of types option | IsType
-val understand_gen_ltac :
- evar_map -> env -> var_map * (identifier * identifier option) list
- -> expected_type:(constr option) -> rawconstr -> constr
-
-(* Generic call to the interpreter from rawconstr to constr, turning
- unresolved holes into metas. Returns also the typing context of
- these metas. Work as [understand_gen] for the rest. *)
-val understand_gen_tcc :
- evar_map -> env -> var_map
- -> constr option -> rawconstr -> open_constr
-
-(* Standard call to get a constr from a rawconstr, resolving implicit args *)
-val understand : evar_map -> env -> rawconstr -> constr
+type var_map = (identifier * unsafe_judgment) list
+type unbound_ltac_var_map = (identifier * identifier option) list
+
+module type S =
+sig
+
+ module Cases : Cases.S
+
+ (* Allow references to syntaxically inexistent variables (i.e., if applied on an inductive) *)
+ val allow_anonymous_refs : bool ref
+
+ (* Generic call to the interpreter from rawconstr to open_constr, leaving
+ unresolved holes as evars and returning the typing contexts of
+ these evars. Work as [understand_gen] for the rest. *)
+
+ val understand_tcc :
+ evar_map -> env -> ?expected_type:types -> rawconstr -> open_constr
+
+ val understand_tcc_evars :
+ evar_defs ref -> env -> typing_constraint -> rawconstr -> constr
+
+ (* More general entry point with evars from ltac *)
+
+ (* Generic call to the interpreter from rawconstr to constr, failing
+ unresolved holes in the rawterm cannot be instantiated.
+
+ In [understand_ltac sigma env ltac_env constraint c],
+
+ sigma : initial set of existential variables (typically dependent subgoals)
+ ltac_env : partial substitution of variables (used for the tactic language)
+ constraint : tell if interpreted as a possibly constrained term or a type
+ *)
+
+ val understand_ltac :
+ evar_map -> env -> var_map * unbound_ltac_var_map ->
+ typing_constraint -> rawconstr -> evar_defs * constr
+
+ (* Standard call to get a constr from a rawconstr, resolving implicit args *)
+
+ val understand : evar_map -> env -> ?expected_type:Term.types ->
+ rawconstr -> constr
+
+ (* Idem but the rawconstr is intended to be a type *)
+
+ val understand_type : evar_map -> env -> rawconstr -> constr
+
+ (* A generalization of the two previous case *)
+
+ val understand_gen : typing_constraint -> evar_map -> env ->
+ rawconstr -> constr
+
+ (* Idem but returns the judgment of the understood term *)
+
+ val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
+
+ (* Idem but do not fail on unresolved evars *)
+ val understand_judgment_tcc : evar_defs ref -> env -> rawconstr -> unsafe_judgment
+
+
+ (*i*)
+ (* Internal of Pretyping...
+ *)
+ val pretype :
+ type_constraint -> env -> evar_defs ref ->
+ var_map * (identifier * identifier option) list ->
+ rawconstr -> unsafe_judgment
+
+ val pretype_type :
+ val_constraint -> env -> evar_defs ref ->
+ var_map * (identifier * identifier option) list ->
+ rawconstr -> unsafe_type_judgment
-(* Idem but the rawconstr is intended to be a type *)
-val understand_type : evar_map -> env -> rawconstr -> constr
+ val pretype_gen :
+ evar_defs ref -> env ->
+ var_map * (identifier * identifier option) list ->
+ typing_constraint -> rawconstr -> constr
-(* Idem but returns the judgment of the understood term *)
-val understand_judgment : evar_map -> env -> rawconstr -> unsafe_judgment
+ (*i*)
+
+end
-(* Idem but returns the judgment of the understood type *)
-val understand_type_judgment : evar_map -> env -> rawconstr
- -> unsafe_type_judgment
+module Pretyping_F (C : Coercion.S) : S
+module Default : S
(* To embed constr in rawconstr *)
+
val constr_in : constr -> Dyn.t
val constr_out : Dyn.t -> constr
-(*i*)
-(* Internal of Pretyping...
- * Unused outside, but useful for debugging
- *)
-val pretype :
- type_constraint -> env -> evar_defs ->
- var_map * (identifier * identifier option) list ->
- rawconstr -> unsafe_judgment
-
-val pretype_type :
- val_constraint -> env -> evar_defs ->
- var_map * (identifier * identifier option) list ->
- rawconstr -> unsafe_type_judgment
-(*i*)
-
-val interp_sort : rawsort -> sorts
-
+val interp_sort : rawsort -> sorts
val interp_elimination_sort : rawsort -> sorts_family
+
diff --git a/pretyping/rawterm.ml b/pretyping/rawterm.ml
index ef4a4670..5d177326 100644
--- a/pretyping/rawterm.ml
+++ b/pretyping/rawterm.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: rawterm.ml,v 1.43.2.4 2004/12/29 10:17:10 herbelin Exp $ *)
+(* $Id: rawterm.ml 8624 2006-03-13 17:38:17Z msozeau $ *)
(*i*)
open Util
@@ -15,6 +15,7 @@ open Sign
open Term
open Libnames
open Nametab
+open Evd
(*i*)
(* Untyped intermediate terms, after ASTs and before constr. *)
@@ -33,8 +34,6 @@ type patvar = identifier
type rawsort = RProp of Term.contents | RType of Univ.universe option
-type fix_kind = RFix of (int array * int) | RCoFix of int
-
type binder_kind = BProd | BLambda | BLetIn
type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier
@@ -48,14 +47,6 @@ type 'a bindings =
type 'a with_bindings = 'a * 'a bindings
-type hole_kind =
- | ImplicitArg of global_reference * (int * identifier option)
- | BinderType of name
- | QuestionMark
- | CasesType
- | InternalHole
- | TomatchTypeParameter of inductive * int
-
type rawconstr =
| RRef of (loc * global_reference)
| RVar of (loc * identifier)
@@ -65,11 +56,9 @@ type rawconstr =
| RLambda of loc * name * rawconstr * rawconstr
| RProd of loc * name * rawconstr * rawconstr
| RLetIn of loc * name * rawconstr * rawconstr
- | RCases of loc * (rawconstr option * rawconstr option ref) *
- (rawconstr * (name * (loc * inductive * name list) option) ref) list *
+ | RCases of loc * rawconstr option *
+ (rawconstr * (name * (loc * inductive * name list) option)) list *
(loc * identifier list * cases_pattern list * rawconstr) list
- | ROrderedCase of loc * case_style * rawconstr option * rawconstr *
- rawconstr array * rawconstr option ref
| RLetTuple of loc * name list * (name * rawconstr option) *
rawconstr * rawconstr
| RIf of loc * rawconstr * (name * rawconstr option) * rawconstr * rawconstr
@@ -77,15 +66,19 @@ type rawconstr =
rawconstr array * rawconstr array
| RSort of loc * rawsort
| RHole of (loc * hole_kind)
- | RCast of loc * rawconstr * rawconstr
+ | RCast of loc * rawconstr * cast_kind * rawconstr
| RDynamic of loc * Dyn.t
and rawdecl = name * rawconstr option * rawconstr
+and fix_recursion_order = RStructRec | RWfRec of rawconstr
+
+and fix_kind = RFix of ((int * fix_recursion_order) array * int) | RCoFix of int
+
let cases_predicate_names tml =
List.flatten (List.map (function
- | (tm,{contents=(na,None)}) -> [na]
- | (tm,{contents=(na,Some (_,_,nal))}) -> na::nal) tml)
+ | (tm,(na,None)) -> [na]
+ | (tm,(na,Some (_,_,nal))) -> na::nal) tml)
(*i - if PRec (_, names, arities, bodies) is in env then arities are
typed in env too and bodies are typed in env enriched by the
@@ -104,12 +97,10 @@ let map_rawconstr f = function
| RLambda (loc,na,ty,c) -> RLambda (loc,na,f ty,f c)
| RProd (loc,na,ty,c) -> RProd (loc,na,f ty,f c)
| RLetIn (loc,na,b,c) -> RLetIn (loc,na,f b,f c)
- | RCases (loc,(tyopt,rtntypopt),tml,pl) ->
- RCases (loc,(option_app f tyopt,ref (option_app f !rtntypopt)),
+ | RCases (loc,rtntypopt,tml,pl) ->
+ RCases (loc,option_app f rtntypopt,
List.map (fun (tm,x) -> (f tm,x)) tml,
List.map (fun (loc,idl,p,c) -> (loc,idl,p,f c)) pl)
- | ROrderedCase (loc,b,tyopt,tm,bv,x) ->
- ROrderedCase (loc,b,option_app f tyopt,f tm, Array.map f bv,ref (option_app f !x))
| RLetTuple (loc,nal,(na,po),b,c) ->
RLetTuple (loc,nal,(na,option_app f po),f b,f c)
| RIf (loc,c,(na,po),b1,b2) ->
@@ -117,7 +108,7 @@ let map_rawconstr f = function
| RRec (loc,fk,idl,bl,tyl,bv) ->
RRec (loc,fk,idl,Array.map (List.map (map_rawdecl f)) bl,
Array.map f tyl,Array.map f bv)
- | RCast (loc,c,t) -> RCast (loc,f c,f t)
+ | RCast (loc,c,k,t) -> RCast (loc,f c,k,f t)
| (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) as x -> x
@@ -147,8 +138,6 @@ let map_rawconstr_with_binders_loc loc g f e = function
let h (_,idl,p,c) = (loc,idl,p,f (List.fold_right g' idl e) c) in
RCases
(loc,option_app (f e) tyopt,List.map (f e) tml, List.map h pl)
- | ROrderedCase (_,b,tyopt,tm,bv) ->
- ROrderedCase (loc,b,option_app (f e) tyopt,f e tm,Array.map (f e) bv)
| RRec (_,fk,idl,tyl,bv) ->
let idl',e' = fold_ident g idl e in
RRec (loc,fk,idl',Array.map (f e) tyl,Array.map (f e') bv)
@@ -168,12 +157,10 @@ let occur_rawconstr id =
| RLambda (loc,na,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
| RProd (loc,na,ty,c) -> (occur ty) or ((na <> Name id) & (occur c))
| RLetIn (loc,na,b,c) -> (occur b) or ((na <> Name id) & (occur c))
- | RCases (loc,(tyopt,rtntypopt),tml,pl) ->
- (occur_option tyopt) or (occur_option !rtntypopt)
+ | RCases (loc,rtntypopt,tml,pl) ->
+ (occur_option rtntypopt)
or (List.exists (fun (tm,_) -> occur tm) tml)
or (List.exists occur_pattern pl)
- | ROrderedCase (loc,b,tyopt,tm,bv,_) ->
- (occur_option tyopt) or (occur tm) or (array_exists occur bv)
| RLetTuple (loc,nal,rtntyp,b,c) ->
occur_return_type rtntyp id
or (occur b) or (not (List.mem (Name id) nal) & (occur c))
@@ -191,7 +178,7 @@ let occur_rawconstr id =
(na=Name id or not(occur_fix bl)) in
occur_fix bl)
idl bl tyl bv)
- | RCast (loc,c,t) -> (occur c) or (occur t)
+ | RCast (loc,c,_,t) -> (occur c) or (occur t)
| (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> false
and occur_pattern (loc,idl,p,c) = not (List.mem id idl) & (occur c)
@@ -202,119 +189,6 @@ let occur_rawconstr id =
in occur
-let rec subst_pat subst pat =
- match pat with
- | PatVar _ -> pat
- | PatCstr (loc,((kn,i),j),cpl,n) ->
- let kn' = subst_kn subst kn
- and cpl' = list_smartmap (subst_pat subst) cpl in
- if kn' == kn && cpl' == cpl then pat else
- PatCstr (loc,((kn',i),j),cpl',n)
-
-let rec subst_raw subst raw =
- match raw with
- | RRef (loc,ref) ->
- let ref' = subst_global subst ref in
- if ref' == ref then raw else
- RRef (loc,ref')
-
- | RVar _ -> raw
- | REvar _ -> raw
- | RPatVar _ -> raw
-
- | RApp (loc,r,rl) ->
- let r' = subst_raw subst r
- and rl' = list_smartmap (subst_raw subst) rl in
- if r' == r && rl' == rl then raw else
- RApp(loc,r',rl')
-
- | RLambda (loc,n,r1,r2) ->
- let r1' = subst_raw subst r1 and r2' = subst_raw subst r2 in
- if r1' == r1 && r2' == r2 then raw else
- RLambda (loc,n,r1',r2')
-
- | RProd (loc,n,r1,r2) ->
- let r1' = subst_raw subst r1 and r2' = subst_raw subst r2 in
- if r1' == r1 && r2' == r2 then raw else
- RProd (loc,n,r1',r2')
-
- | RLetIn (loc,n,r1,r2) ->
- let r1' = subst_raw subst r1 and r2' = subst_raw subst r2 in
- if r1' == r1 && r2' == r2 then raw else
- RLetIn (loc,n,r1',r2')
-
- | RCases (loc,(ro,rtno),rl,branches) ->
- let ro' = option_smartmap (subst_raw subst) ro
- and rtno' = ref (option_smartmap (subst_raw subst) !rtno)
- and rl' = list_smartmap (fun (a,x as y) ->
- let a' = subst_raw subst a in
- let (n,topt) = !x in
- let topt' = option_smartmap
- (fun (loc,(sp,i),x as t) ->
- let sp' = subst_kn subst sp in
- if sp == sp' then t else (loc,(sp',i),x)) topt in
- if a == a' && topt == topt' then y else (a',ref (n,topt'))) rl
- and branches' = list_smartmap
- (fun (loc,idl,cpl,r as branch) ->
- let cpl' = list_smartmap (subst_pat subst) cpl
- and r' = subst_raw subst r in
- if cpl' == cpl && r' == r then branch else
- (loc,idl,cpl',r'))
- branches
- in
- if ro' == ro && rl' == rl && branches' == branches then raw else
- RCases (loc,(ro',rtno'),rl',branches')
-
- | ROrderedCase (loc,b,ro,r,ra,x) ->
- let ro' = option_smartmap (subst_raw subst) ro
- and r' = subst_raw subst r
- and ra' = array_smartmap (subst_raw subst) ra in
- if ro' == ro && r' == r && ra' == ra then raw else
- ROrderedCase (loc,b,ro',r',ra',x)
-
- | RLetTuple (loc,nal,(na,po),b,c) ->
- let po' = option_smartmap (subst_raw subst) po
- and b' = subst_raw subst b
- and c' = subst_raw subst c in
- if po' == po && b' == b && c' == c then raw else
- RLetTuple (loc,nal,(na,po'),b',c')
-
- | RIf (loc,c,(na,po),b1,b2) ->
- let po' = option_smartmap (subst_raw subst) po
- and b1' = subst_raw subst b1
- and b2' = subst_raw subst b2
- and c' = subst_raw subst c in
- if c' == c & po' == po && b1' == b1 && b2' == b2 then raw else
- RIf (loc,c',(na,po'),b1',b2')
-
- | RRec (loc,fix,ida,bl,ra1,ra2) ->
- let ra1' = array_smartmap (subst_raw subst) ra1
- and ra2' = array_smartmap (subst_raw subst) ra2 in
- let bl' = array_smartmap
- (list_smartmap (fun (na,obd,ty as dcl) ->
- let ty' = subst_raw subst ty in
- let obd' = option_smartmap (subst_raw subst) obd in
- if ty'==ty & obd'==obd then dcl else (na,obd',ty')))
- bl in
- if ra1' == ra1 && ra2' == ra2 && bl'==bl then raw else
- RRec (loc,fix,ida,bl',ra1',ra2')
-
- | RSort _ -> raw
-
- | RHole (loc,ImplicitArg (ref,i)) ->
- let ref' = subst_global subst ref in
- if ref' == ref then raw else
- RHole (loc,ImplicitArg (ref',i))
- | RHole (loc, (BinderType _ | QuestionMark | CasesType |
- InternalHole | TomatchTypeParameter _)) -> raw
-
- | RCast (loc,r1,r2) ->
- let r1' = subst_raw subst r1 and r2' = subst_raw subst r2 in
- if r1' == r1 && r2' == r2 then raw else
- RCast (loc,r1',r2')
-
- | RDynamic _ -> raw
-
let loc_of_rawconstr = function
| RRef (loc,_) -> loc
| RVar (loc,_) -> loc
@@ -325,15 +199,47 @@ let loc_of_rawconstr = function
| RProd (loc,_,_,_) -> loc
| RLetIn (loc,_,_,_) -> loc
| RCases (loc,_,_,_) -> loc
- | ROrderedCase (loc,_,_,_,_,_) -> loc
| RLetTuple (loc,_,_,_,_) -> loc
| RIf (loc,_,_,_,_) -> loc
| RRec (loc,_,_,_,_,_) -> loc
| RSort (loc,_) -> loc
| RHole (loc,_) -> loc
- | RCast (loc,_,_) -> loc
+ | RCast (loc,_,_,_) -> loc
| RDynamic (loc,_) -> loc
+(**********************************************************************)
+(* Conversion from rawconstr to cases pattern, if possible *)
+
+let rec cases_pattern_of_rawconstr na = function
+ | RVar (loc,id) when na<>Anonymous ->
+ (* Unable to manage the presence of both an alias and a variable *)
+ raise Not_found
+ | RVar (loc,id) -> PatVar (loc,Name id)
+ | RHole (loc,_) -> PatVar (loc,na)
+ | RRef (loc,ConstructRef cstr) ->
+ PatCstr (loc,cstr,[],na)
+ | RApp (loc,RRef (_,ConstructRef cstr),l) ->
+ PatCstr (loc,cstr,List.map (cases_pattern_of_rawconstr Anonymous) l,na)
+ | _ -> raise Not_found
+
+(* Turn a closed cases pattern into a rawconstr *)
+let rec rawconstr_of_closed_cases_pattern_aux = function
+ | PatCstr (loc,cstr,[],Anonymous) ->
+ RRef (loc,ConstructRef cstr)
+ | PatCstr (loc,cstr,l,Anonymous) ->
+ let ref = RRef (loc,ConstructRef cstr) in
+ RApp (loc,ref, List.map rawconstr_of_closed_cases_pattern_aux l)
+ | _ -> raise Not_found
+
+let rawconstr_of_closed_cases_pattern = function
+ | PatCstr (loc,cstr,l,na) ->
+ na,rawconstr_of_closed_cases_pattern_aux (PatCstr (loc,cstr,l,Anonymous))
+ | _ ->
+ raise Not_found
+
+(**********************************************************************)
+(* Reduction expressions *)
+
type 'a raw_red_flag = {
rBeta : bool;
rIota : bool;
@@ -357,6 +263,7 @@ type ('a,'b) red_expr_gen =
| Fold of 'a list
| Pattern of 'a occurrences list
| ExtraRedExpr of string
+ | CbvVm
type ('a,'b) may_eval =
| ConstrTerm of 'a
diff --git a/pretyping/rawterm.mli b/pretyping/rawterm.mli
index 97e11af6..22317b5f 100644
--- a/pretyping/rawterm.mli
+++ b/pretyping/rawterm.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: rawterm.mli,v 1.47.2.5 2005/01/21 16:42:37 herbelin Exp $ i*)
+(*i $Id: rawterm.mli 8624 2006-03-13 17:38:17Z msozeau $ i*)
(*i*)
open Util
@@ -31,8 +31,6 @@ type patvar = identifier
type rawsort = RProp of Term.contents | RType of Univ.universe option
-type fix_kind = RFix of (int array * int) | RCoFix of int
-
type binder_kind = BProd | BLambda | BLetIn
type quantified_hypothesis = AnonHyp of int | NamedHyp of identifier
@@ -46,14 +44,6 @@ type 'a bindings =
type 'a with_bindings = 'a * 'a bindings
-type hole_kind =
- | ImplicitArg of global_reference * (int * identifier option)
- | BinderType of name
- | QuestionMark
- | CasesType
- | InternalHole
- | TomatchTypeParameter of inductive * int
-
type rawconstr =
| RRef of (loc * global_reference)
| RVar of (loc * identifier)
@@ -63,26 +53,27 @@ type rawconstr =
| RLambda of loc * name * rawconstr * rawconstr
| RProd of loc * name * rawconstr * rawconstr
| RLetIn of loc * name * rawconstr * rawconstr
- | RCases of loc * (rawconstr option * rawconstr option ref) *
- (rawconstr * (name * (loc * inductive * name list) option) ref) list *
+ | RCases of loc * rawconstr option *
+ (rawconstr * (name * (loc * inductive * name list) option)) list *
(loc * identifier list * cases_pattern list * rawconstr) list
- | ROrderedCase of loc * case_style * rawconstr option * rawconstr *
- rawconstr array * rawconstr option ref
| RLetTuple of loc * name list * (name * rawconstr option) *
rawconstr * rawconstr
| RIf of loc * rawconstr * (name * rawconstr option) * rawconstr * rawconstr
| RRec of loc * fix_kind * identifier array * rawdecl list array *
rawconstr array * rawconstr array
| RSort of loc * rawsort
- | RHole of (loc * hole_kind)
- | RCast of loc * rawconstr * rawconstr
+ | RHole of (loc * Evd.hole_kind)
+ | RCast of loc * rawconstr * cast_kind * rawconstr
| RDynamic of loc * Dyn.t
and rawdecl = name * rawconstr option * rawconstr
+and fix_recursion_order = RStructRec | RWfRec of rawconstr
+
+and fix_kind = RFix of ((int * fix_recursion_order) array * int) | RCoFix of int
+
val cases_predicate_names :
- (rawconstr * (name * (loc * inductive * name list) option) ref) list ->
- name list
+ (rawconstr * (name * (loc * inductive * name list) option)) list -> name list
(*i - if PRec (_, names, arities, bodies) is in env then arities are
typed in env too and bodies are typed in env enriched by the
@@ -107,7 +98,18 @@ val occur_rawconstr : identifier -> rawconstr -> bool
val loc_of_rawconstr : rawconstr -> loc
-val subst_raw : Names.substitution -> rawconstr -> rawconstr
+(**********************************************************************)
+(* Conversion from rawconstr to cases pattern, if possible *)
+
+(* Take the current alias as parameter, raise Not_found if *)
+(* translation is impossible *)
+
+val cases_pattern_of_rawconstr : name -> rawconstr -> cases_pattern
+
+val rawconstr_of_closed_cases_pattern : cases_pattern -> name * rawconstr
+
+(**********************************************************************)
+(* Reduction expressions *)
type 'a raw_red_flag = {
rBeta : bool;
@@ -131,6 +133,7 @@ type ('a,'b) red_expr_gen =
| Fold of 'a list
| Pattern of 'a occurrences list
| ExtraRedExpr of string
+ | CbvVm
type ('a,'b) may_eval =
| ConstrTerm of 'a
diff --git a/pretyping/recordops.ml b/pretyping/recordops.ml
index 3e73cfee..87997d99 100755..100644
--- a/pretyping/recordops.ml
+++ b/pretyping/recordops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: recordops.ml,v 1.26.2.2 2005/11/29 21:40:52 letouzey Exp $ *)
+(* $Id: recordops.ml 8642 2006-03-17 10:09:02Z notin $ *)
open Util
open Pp
@@ -19,133 +19,199 @@ open Typeops
open Libobject
open Library
open Classops
+open Mod_subst
-(*s Une structure S est un type inductif non récursif à un seul
- constructeur (de nom par défaut Build_S) *)
+(*s A structure S is a non recursive inductive type with a single
+ constructor (the name of which defaults to Build_S) *)
(* Table des structures: le nom de la structure (un [inductive]) donne
le nom du constructeur, le nombre de paramètres et pour chaque
- argument réels du constructeur, le noms de la projection
- correspondante, si valide *)
+ argument réel du constructeur, le nom de la projection
+ correspondante, si valide, et un booléen disant si c'est une vraie
+ projection ou bien une fonction constante (associée à un LetIn) *)
type struc_typ = {
s_CONST : identifier;
s_PARAM : int;
+ s_PROJKIND : bool list;
s_PROJ : constant option list }
let structure_table = ref (Indmap.empty : struc_typ Indmap.t)
-let projection_table = ref KNmap.empty
+let projection_table = ref Cmap.empty
let option_fold_right f p e = match p with Some a -> f a e | None -> e
-let cache_structure (_,(ind,struc)) =
+let load_structure i (_,(ind,id,kl,projs)) =
+ let n = (fst (Global.lookup_inductive ind)).Declarations.mind_nparams in
+ let struc =
+ { s_CONST = id; s_PARAM = n; s_PROJ = projs; s_PROJKIND = kl } in
structure_table := Indmap.add ind struc !structure_table;
projection_table :=
- List.fold_right (option_fold_right (fun proj -> KNmap.add proj struc))
- struc.s_PROJ !projection_table
+ List.fold_right (option_fold_right (fun proj -> Cmap.add proj struc))
+ projs !projection_table
-let subst_structure (_,subst,((kn,i),struc as obj)) =
+let cache_structure o =
+ load_structure 1 o
+
+let subst_structure (_,subst,((kn,i),id,kl,projs as obj)) =
let kn' = subst_kn subst kn in
- let proj' = list_smartmap
- (option_smartmap (subst_kn subst))
- struc.s_PROJ
+ let projs' =
+ (* invariant: struc.s_PROJ is an evaluable reference. Thus we can take *)
+ (* the first component of subst_con. *)
+ list_smartmap
+ (option_smartmap (fun kn -> fst (subst_con subst kn)))
+ projs
in
- if proj' == struc.s_PROJ && kn' == kn then obj else
- (kn',i),{struc with s_PROJ = proj'}
+ if projs' == projs && kn' == kn then obj else
+ ((kn',i),id,kl,projs')
+
+let discharge_structure (_,(ind,id,kl,projs)) =
+ Some (Lib.discharge_inductive ind, id, kl,
+ List.map (option_app Lib.discharge_con) projs)
let (inStruc,outStruc) =
declare_object {(default_object "STRUCTURE") with
- cache_function = cache_structure;
- load_function = (fun _ o -> cache_structure o);
- subst_function = subst_structure;
- classify_function = (fun (_,x) -> Substitute x);
- export_function = (function x -> Some x) }
+ cache_function = cache_structure;
+ load_function = load_structure;
+ subst_function = subst_structure;
+ classify_function = (fun (_,x) -> Substitute x);
+ discharge_function = discharge_structure;
+ export_function = (function x -> Some x) }
-let add_new_struc (s,c,n,l) =
- Lib.add_anonymous_leaf (inStruc (s,{s_CONST=c;s_PARAM=n;s_PROJ=l}))
+let declare_structure (s,c,_,kl,pl) =
+ Lib.add_anonymous_leaf (inStruc (s,c,kl,pl))
-let find_structure indsp = Indmap.find indsp !structure_table
+let lookup_structure indsp = Indmap.find indsp !structure_table
let find_projection_nparams = function
- | ConstRef cst -> (KNmap.find cst !projection_table).s_PARAM
+ | ConstRef cst -> (Cmap.find cst !projection_table).s_PARAM
| _ -> raise Not_found
-(*s Un "object" est une fonction construisant une instance d'une structure *)
+
+(************************************************************************)
+(*s A canonical structure declares "canonical" conversion hints between *)
+(* the effective components of a structure and the projections of the *)
+(* structure *)
(* Table des definitions "object" : pour chaque object c,
c := [x1:B1]...[xk:Bk](Build_R a1...am t1...t_n)
- avec ti = (ci ui1...uir)
+ If ti has the form (ci ui1...uir) where ci is a global reference and
+if the corresponding projection Li of the structure R is defined, one
+declare a "conversion" between ci and Li
+
+ x1:B1..xk:Bk |- (Li a1..am (c x1..xk)) =_conv (ci ui1...uir)
- Pour tout ci, et Li, la ième projection de la structure R (si
- définie), on déclare une "coercion"
+that maps the pair (Li,ci) to the following data
o_DEF = c
o_TABS = B1...Bk
o_PARAMS = a1...am
o_TCOMP = ui1...uir
+
*)
type obj_typ = {
o_DEF : constr;
- o_TABS : constr list; (* dans l'ordre *)
- o_TPARAMS : constr list; (* dans l'ordre *)
- o_TCOMPS : constr list } (* dans l'ordre *)
-
-let subst_obj subst obj =
- let o_DEF' = subst_mps subst obj.o_DEF in
- let o_TABS' = list_smartmap (subst_mps subst) obj.o_TABS in
- let o_TPARAMS' = list_smartmap (subst_mps subst) obj.o_TPARAMS in
- let o_TCOMPS' = list_smartmap (subst_mps subst) obj.o_TCOMPS in
- if o_DEF' == obj.o_DEF
- && o_TABS' == obj.o_TABS
- && o_TPARAMS' == obj.o_TPARAMS
- && o_TCOMPS' == obj.o_TCOMPS
- then
- obj
- else
- { o_DEF = o_DEF' ;
- o_TABS = o_TABS' ;
- o_TPARAMS = o_TPARAMS' ;
- o_TCOMPS = o_TCOMPS' }
+ o_TABS : constr list; (* ordered *)
+ o_TPARAMS : constr list; (* ordered *)
+ o_TCOMPS : constr list } (* ordered *)
let object_table =
(ref [] : ((global_reference * global_reference) * obj_typ) list ref)
-let cache_object (_,x) = object_table := x :: (!object_table)
-
-let subst_object (_,subst,((r1,r2),o as obj)) =
- let r1' = subst_global subst r1 in
- let r2' = subst_global subst r2 in
- let o' = subst_obj subst o in
- if r1' == r1 && r2' == r2 && o' == o then obj else
- (r1',r2'),o'
-
-let (inObjDef,outObjDef) =
- declare_object {(default_object "OBJDEF") with
- open_function = (fun i o -> if i=1 then cache_object o);
- cache_function = cache_object;
- subst_function = subst_object;
- classify_function = (fun (_,x) -> Substitute x);
- export_function = (function x -> Some x) }
-
-let add_new_objdef (o,c,la,lp,l) =
- try
- let _ = List.assoc o !object_table in ()
- with Not_found ->
- Lib.add_anonymous_leaf
- (inObjDef (o,{o_DEF=c;o_TABS=la;o_TPARAMS=lp;o_TCOMPS=l}))
-
-let cache_objdef1 (_,sp) = ()
-
-let (inObjDef1,outObjDef1) =
- declare_object {(default_object "OBJDEF1") with
- open_function = (fun i o -> if i=1 then cache_objdef1 o);
- cache_function = cache_objdef1;
- export_function = (function x -> Some x) }
-
-let objdef_info o = List.assoc o !object_table
+let canonical_projections () = !object_table
+
+let keep_true_projections projs kinds =
+ map_succeed (function (p,true) -> p | _ -> failwith "")
+ (List.combine projs kinds)
+
+(* Intended to always success *)
+let compute_canonical_projections (con,ind) =
+ let v = mkConst con in
+ let c = Environ.constant_value (Global.env()) con in
+ let lt,t = Reductionops.splay_lambda (Global.env()) Evd.empty c in
+ let lt = List.rev (List.map snd lt) in
+ let args = snd (decompose_app t) in
+ let { s_PARAM = p; s_PROJ = lpj; s_PROJKIND = kl } = lookup_structure ind in
+ let params, projs = list_chop p args in
+ let lpj = keep_true_projections lpj kl in
+ let lps = List.combine lpj projs in
+ let comp =
+ List.fold_left
+ (fun l (spopt,t) -> (* comp=components *)
+ match spopt with
+ | Some proji_sp ->
+ let c, args = decompose_app t in
+ (try (ConstRef proji_sp, global_of_constr c, args) :: l
+ with Not_found -> l)
+ | _ -> l)
+ [] lps in
+ List.map (fun (refi,c,argj) ->
+ (refi,c),{o_DEF=v; o_TABS=lt; o_TPARAMS=params; o_TCOMPS=argj})
+ comp
+
+let open_canonical_structure i (_,o) =
+ if i=1 then
+ let lo = compute_canonical_projections o in
+ List.iter (fun (o,_ as x) ->
+ if not (List.mem_assoc o !object_table) then
+ object_table := x :: (!object_table)) lo
+
+let cache_canonical_structure o =
+ open_canonical_structure 1 o
+
+let subst_canonical_structure (_,subst,(cst,ind as obj)) =
+ (* invariant: cst is an evaluable reference. Thus we can take *)
+ (* the first component of subst_con. *)
+ let cst' = fst (subst_con subst cst) in
+ let ind' = Inductiveops.subst_inductive subst ind in
+ if cst' == cst & ind' == ind then obj else (cst',ind')
+
+let discharge_canonical_structure (_,(cst,ind)) =
+ Some (Lib.discharge_con cst,Lib.discharge_inductive ind)
+
+let (inCanonStruc,outCanonStruct) =
+ declare_object {(default_object "CANONICAL-STRUCTURE") with
+ open_function = open_canonical_structure;
+ cache_function = cache_canonical_structure;
+ subst_function = subst_canonical_structure;
+ classify_function = (fun (_,x) -> Substitute x);
+ discharge_function = discharge_canonical_structure;
+ export_function = (function x -> Some x) }
+
+let add_canonical_structure x = Lib.add_anonymous_leaf (inCanonStruc x)
+
+(*s High-level declaration of a canonical structure *)
+
+let error_not_structure ref =
+ errorlabstrm "object_declare"
+ (Nameops.pr_id (id_of_global ref) ++ str" is not a structure object")
+
+let check_and_decompose_canonical_structure ref =
+ let sp = match ref with ConstRef sp -> sp | _ -> error_not_structure ref in
+ let env = Global.env () in
+ let vc = match Environ.constant_opt_value env sp with
+ | Some vc -> vc
+ | None -> error_not_structure ref in
+ let f,args = match kind_of_term (snd (decompose_lam vc)) with
+ | App (f,args) -> f,args
+ | _ -> error_not_structure ref in
+ let indsp = match kind_of_term f with
+ | Construct (indsp,1) -> indsp
+ | _ -> error_not_structure ref in
+ let s = try lookup_structure indsp with Not_found -> error_not_structure ref in
+ if s.s_PARAM + List.length s.s_PROJ > Array.length args then
+ error_not_structure ref;
+ (sp,indsp)
+
+let declare_canonical_structure ref =
+ add_canonical_structure (check_and_decompose_canonical_structure ref)
+
+let outCanonicalStructure x = fst (outCanonStruct x)
+
+let lookup_canonical_conversion o = List.assoc o !object_table
let freeze () =
!structure_table, !projection_table, !object_table
@@ -154,7 +220,7 @@ let unfreeze (s,p,o) =
structure_table := s; projection_table := p; object_table := o
let init () =
- structure_table := Indmap.empty; projection_table := KNmap.empty;
+ structure_table := Indmap.empty; projection_table := Cmap.empty;
object_table:=[]
let _ = init()
diff --git a/pretyping/recordops.mli b/pretyping/recordops.mli
index a458b7b3..1e061dc6 100755
--- a/pretyping/recordops.mli
+++ b/pretyping/recordops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: recordops.mli,v 1.15.2.2 2005/11/29 21:40:52 letouzey Exp $ i*)
+(*i $Id: recordops.mli 6748 2005-02-18 22:17:50Z herbelin $ i*)
(*i*)
open Names
@@ -18,34 +18,36 @@ open Libobject
open Library
(*i*)
+(*s A structure S is a non recursive inductive type with a single
+ constructor (the name of which defaults to Build_S) *)
+
type struc_typ = {
s_CONST : identifier;
s_PARAM : int;
+ s_PROJKIND : bool list;
s_PROJ : constant option list }
-val add_new_struc :
- inductive * identifier * int * constant option list -> unit
+val declare_structure :
+ inductive * identifier * int * bool list * constant option list -> unit
-(* [find_structure isp] returns the infos associated to inductive path
+(* [lookup_structure isp] returns the infos associated to inductive path
[isp] if it corresponds to a structure, otherwise fails with [Not_found] *)
-val find_structure : inductive -> struc_typ
+val lookup_structure : inductive -> struc_typ
(* raise [Not_found] if not a projection *)
val find_projection_nparams : global_reference -> int
+(*s A canonical structure declares "canonical" conversion hints between *)
+(* the effective components of a structure and the projections of the *)
+(* structure *)
+
type obj_typ = {
o_DEF : constr;
- o_TABS : constr list; (* dans l'ordre *)
- o_TPARAMS : constr list; (* dans l'ordre *)
- o_TCOMPS : constr list } (* dans l'ordre *)
+ o_TABS : constr list; (* ordered *)
+ o_TPARAMS : constr list; (* ordered *)
+ o_TCOMPS : constr list } (* ordered *)
-val objdef_info : (global_reference * global_reference) -> obj_typ
-val add_new_objdef :
- (global_reference * global_reference) * Term.constr * Term.constr list *
- Term.constr list * Term.constr list -> unit
-
-
-val inStruc : inductive * struc_typ -> obj
-val outStruc : obj -> inductive * struc_typ
-val inObjDef1 : kernel_name -> obj
-val outObjDef1 : obj -> kernel_name
+val lookup_canonical_conversion : (global_reference * global_reference) -> obj_typ
+val declare_canonical_structure : global_reference -> unit
+val canonical_projections : unit ->
+ ((global_reference * global_reference) * obj_typ) list
diff --git a/pretyping/reductionops.ml b/pretyping/reductionops.ml
index a030dcf2..b590f743 100644
--- a/pretyping/reductionops.ml
+++ b/pretyping/reductionops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: reductionops.ml,v 1.6.2.2 2004/07/16 19:30:46 herbelin Exp $ *)
+(* $Id: reductionops.ml 8708 2006-04-14 08:13:02Z jforest $ *)
open Pp
open Util
@@ -17,7 +17,6 @@ open Univ
open Evd
open Declarations
open Environ
-open Instantiate
open Closure
open Esubst
open Reduction
@@ -48,7 +47,7 @@ type local_state_reduction_function = state -> state
let rec whd_state (x, stack as s) =
match kind_of_term x with
| App (f,cl) -> whd_state (f, append_stack cl stack)
- | Cast (c,_) -> whd_state (c, stack)
+ | Cast (c,_,_) -> whd_state (c, stack)
| _ -> s
let appterm_of_stack (f,s) = (f,list_of_stack s)
@@ -189,7 +188,7 @@ let contract_cofix (bodynum,(types,names,bodies as typedbodies)) =
let reduce_mind_case mia =
match kind_of_term mia.mconstr with
- | Construct (ind_sp,i as cstr_sp) ->
+ | Construct (ind_sp,i) ->
(* let ncargs = (fst mia.mci).(i-1) in*)
let real_cargs = list_skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1),real_cargs)
@@ -255,7 +254,7 @@ let rec whd_state_gen flags env sigma =
| Some body -> whrec (body, stack)
| None -> s)
| LetIn (_,b,_,c) when red_zeta flags -> stacklam whrec [b] c stack
- | Cast (c,_) -> whrec (c, stack)
+ | Cast (c,_,_) -> whrec (c, stack)
| App (f,cl) -> whrec (f, append_stack cl stack)
| Lambda (na,t,c) ->
(match decomp_stack stack with
@@ -300,7 +299,7 @@ let local_whd_state_gen flags =
let rec whrec (x, stack as s) =
match kind_of_term x with
| LetIn (_,b,_,c) when red_zeta flags -> stacklam whrec [b] c stack
- | Cast (c,_) -> whrec (c, stack)
+ | Cast (c,_,_) -> whrec (c, stack)
| App (f,cl) -> whrec (f, append_stack cl stack)
| Lambda (_,_,c) ->
(match decomp_stack stack with
@@ -339,6 +338,7 @@ let local_whd_state_gen flags =
in
whrec
+
(* 1. Beta Reduction Functions *)
let whd_beta_state = local_whd_state_gen beta
@@ -427,12 +427,14 @@ let whd_betadeltaiota_nolet env sigma x =
(* Replacing defined evars for error messages *)
let rec whd_evar sigma c =
match kind_of_term c with
- | Evar (ev,args) when Evd.in_dom sigma ev & Evd.is_defined sigma ev ->
- whd_evar sigma (Instantiate.existential_value sigma (ev,args))
+ | Evar (ev,args)
+ when Evd.in_dom sigma ev & Evd.is_defined sigma ev ->
+ whd_evar sigma (Evd.existential_value sigma (ev,args))
+ | Sort s when is_sort_variable sigma s -> whd_sort_variable sigma c
| _ -> collapse_appl c
-let nf_evar sigma =
- local_strong (whd_evar sigma)
+let nf_evar evd =
+ local_strong (whd_evar evd)
(* lazy reduction functions. The infos must be created for each term *)
let clos_norm_flags flgs env sigma t =
@@ -443,6 +445,120 @@ let nf_betaiota = clos_norm_flags Closure.betaiota empty_env Evd.empty
let nf_betadeltaiota env sigma =
clos_norm_flags Closure.betadeltaiota env sigma
+
+(* Attention reduire un beta-redexe avec un argument qui n'est pas
+ une variable, peut changer enormement le temps de conversion lors
+ du type checking :
+ (fun x => x + x) M
+*)
+let nf_betaiotaevar_preserving_vm_cast env sigma t =
+ let push decl (env,subst) =
+ (Environ.push_rel decl env, Esubst.subs_lift subst) in
+ let cons decl v (env, subst) = (push_rel decl env, Esubst.subs_cons (v,subst)) in
+
+ let app_stack t (f, stack) =
+ let t' = app_stack (f,stack) in
+ match kind_of_term t, kind_of_term t' with
+ | App(f,args), App(f',args') when f == f' && array_for_all2 (==) args args' -> t
+ | _ -> t'
+ in
+ let rec whrec (env, subst as es) (t, stack as s) =
+ match kind_of_term t with
+ | Rel i ->
+ let t' =
+ match Esubst.expand_rel i subst with
+ | Inl (k,e) -> lift k e
+ | Inr (k,None) -> mkRel k
+ | Inr (k, Some p) -> lift (k-p) (mkRel p) (*??????? == mkRel k ! Julien *)
+ (* Est correct ??? *)
+ in
+ if t = t' then s else t', stack
+ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _ -> s
+ | Evar (e,al) ->
+ let al' = Array.map (norm es) al in
+ begin match existential_opt_value sigma (e,al') with
+ | Some body -> whrec (env,Esubst.ESID 0) (body, stack) (**** ????? ****)
+ | None ->
+ if array_for_all2 (==) al al' then s else (mkEvar (e, al'), stack)
+ end
+ | Cast (c,VMcast,t) ->
+ let c' = norm es c in
+ let t' = norm es t in
+ if c == c' && t == t' then s
+ else (mkCast(c',VMcast,t'),stack)
+ | Cast (c,DEFAULTcast,_) ->
+ whrec es (c, stack)
+
+ | Prod (na,t,c) ->
+ let t' = norm es t in
+ let c' = norm (push (na, None, t') es) c in
+ if t==t' && c==c' then s else (mkProd (na, t', c'), stack)
+
+ | Lambda (na,t,c) ->
+ begin match decomp_stack stack with
+ | Some (a,m) ->
+ begin match kind_of_term a with
+ | Rel i when not (evaluable_rel i env) ->
+ whrec (cons (na,None,t) a es) (c, m)
+ | Var id when not (evaluable_named id env)->
+ whrec (cons (na,None,t) a es) (c, m)
+ | _ ->
+ let t' = norm es t in
+ let c' = norm (push (na, None, t') es) c in
+ if t == t' && c == c' then s
+ else mkLambda (na, t', c'), stack
+ end
+ | _ ->
+ let t' = norm es t in
+ let c' = norm (push (na, None, t') es) c in
+ if t == t' && c == c' then s
+ else mkLambda(na,t',c'),stack
+
+ end
+ | LetIn (na,b,t,c) ->
+ let b' = norm es b in
+ let t' = norm es t in
+ let c' = norm (push (na, Some b', t') es) c in
+ if b==b' && t==t' && c==c' then s
+ else mkLetIn (na, b', t', c'), stack
+
+ | App (f,cl) ->
+ let cl' = Array.map (norm es) cl in
+ whrec es (f, append_stack cl' stack)
+
+ | Case (ci,p,d,lf) ->
+ let (c,cargs) = whrec es (d, empty_stack) in
+ if reducible_mind_case c then
+ whrec es (reduce_mind_case
+ {mP=p; mconstr=c; mcargs=list_of_stack cargs;
+ mci=ci; mlf=lf}, stack)
+ else
+ let p' = norm es p in
+ let d' = app_stack d (c,cargs) in
+ let lf' = Array.map (norm es) lf in
+ if p==p' && d==d' && array_for_all2 (==) lf lf' then s
+ else (mkCase (ci, p', d', lf'), stack)
+ | Fix (ln,(lna,tl,bl)) ->
+ let tl' = Array.map (norm es) tl in
+ let es' =
+ array_fold_left2 (fun es na t -> push (na,None,t) es) es lna tl' in
+ let bl' = Array.map (norm es') bl in
+ if array_for_all2 (==) tl tl' && array_for_all2 (==) bl bl'
+ then s
+ else (mkFix (ln,(lna,tl',bl')), stack)
+ | CoFix(ln,(lna,tl,bl)) ->
+ let tl' = Array.map (norm es) tl in
+ let es' =
+ array_fold_left2 (fun es na t -> push (na,None,t) es) es lna tl in
+ let bl' = Array.map (norm es') bl in
+ if array_for_all2 (==) tl tl' && array_for_all2 (==) bl bl'
+ then s
+ else (mkCoFix (ln,(lna,tl',bl')), stack)
+
+ and norm es t = app_stack t (whrec es (t,empty_stack)) in
+ norm (env, Esubst.ESID 0) t
+
+
(* lazy weak head reduction functions *)
let whd_flags flgs env sigma t =
whd_val (create_clos_infos flgs env) (inject (nf_evar sigma t))
@@ -462,10 +578,6 @@ let fhnf_apply info k h a = Profile.profile4 fakey fhnf_apply info k h a;;
type conversion_test = constraints -> constraints
-type conv_pb =
- | CONV
- | CUMUL
-
let pb_is_equal pb = pb = CONV
let pb_equal = function
@@ -515,22 +627,22 @@ let plain_instance s c =
let rec irec u = match kind_of_term u with
| Meta p -> (try List.assoc p s with Not_found -> u)
| App (f,l) when isCast f ->
- let (f,t) = destCast f in
+ let (f,_,t) = destCast f in
let l' = Array.map irec l in
(match kind_of_term f with
- | Meta p ->
- (* Don't flatten application nodes: this is used to extract a
- proof-term from a proof-tree and we want to keep the structure
- of the proof-tree *)
- (try let g = List.assoc p s in
- match kind_of_term g with
- | App _ ->
- let h = id_of_string "H" in
- mkLetIn (Name h,g,t,mkApp(mkRel 1,Array.map (lift 1) l'))
- | _ -> mkApp (g,l')
- with Not_found -> mkApp (f,l'))
- | _ -> mkApp (irec f,l'))
- | Cast (m,_) when isMeta m ->
+ | Meta p ->
+ (* Don't flatten application nodes: this is used to extract a
+ proof-term from a proof-tree and we want to keep the structure
+ of the proof-tree *)
+ (try let g = List.assoc p s in
+ match kind_of_term g with
+ | App _ ->
+ let h = id_of_string "H" in
+ mkLetIn (Name h,g,t,mkApp(mkRel 1,Array.map (lift 1) l'))
+ | _ -> mkApp (g,l')
+ with Not_found -> mkApp (f,l'))
+ | _ -> mkApp (irec f,l'))
+ | Cast (m,_,_) when isMeta m ->
(try List.assoc (destMeta m) s with Not_found -> u)
| _ -> map_constr irec u
in
@@ -580,17 +692,28 @@ let splay_prod env sigma =
in
decrec env []
+let splay_lambda env sigma =
+ let rec decrec env m c =
+ let t = whd_betadeltaiota env sigma c in
+ match kind_of_term t with
+ | Lambda (n,a,c0) ->
+ decrec (push_rel (n,None,a) env)
+ ((n,a)::m) c0
+ | _ -> m,t
+ in
+ decrec env []
+
let splay_prod_assum env sigma =
let rec prodec_rec env l c =
let t = whd_betadeltaiota_nolet env sigma c in
- match kind_of_term c with
+ match kind_of_term t with
| Prod (x,t,c) ->
prodec_rec (push_rel (x,None,t) env)
(Sign.add_rel_decl (x, None, t) l) c
| LetIn (x,b,t,c) ->
prodec_rec (push_rel (x, Some b, t) env)
(Sign.add_rel_decl (x, Some b, t) l) c
- | Cast (c,_) -> prodec_rec env l c
+ | Cast (c,_,_) -> prodec_rec env l c
| _ -> l,t
in
prodec_rec env Sign.empty_rel_context
@@ -613,6 +736,13 @@ let decomp_n_prod env sigma n =
in
decrec env n Sign.empty_rel_context
+exception NotASort
+
+let decomp_sort env sigma t =
+ match kind_of_term (whd_betadeltaiota env sigma t) with
+ | Sort s -> s
+ | _ -> raise NotASort
+
(* One step of approximation *)
let rec apprec env sigma s =
@@ -715,3 +845,27 @@ let is_info_type env sigma t =
(s = Prop Pos) ||
(s <> Prop Null &&
try info_arity env sigma t.utj_val with IsType -> true)
+
+(*************************************)
+(* Metas *)
+
+let meta_value evd mv =
+ let rec valrec mv =
+ try
+ let b = meta_fvalue evd mv in
+ instance
+ (List.map (fun mv' -> (mv',valrec mv')) (Metaset.elements b.freemetas))
+ b.rebus
+ with Anomaly _ | Not_found ->
+ mkMeta mv
+ in
+ valrec mv
+
+let meta_instance env b =
+ let c_sigma =
+ List.map
+ (fun mv -> (mv,meta_value env mv)) (Metaset.elements b.freemetas)
+ in
+ instance c_sigma b.rebus
+
+let nf_meta env c = meta_instance env (mk_freelisted c)
diff --git a/pretyping/reductionops.mli b/pretyping/reductionops.mli
index 65cdd5cd..ff55cc0e 100644
--- a/pretyping/reductionops.mli
+++ b/pretyping/reductionops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: reductionops.mli,v 1.8.2.2 2004/07/16 19:30:46 herbelin Exp $ i*)
+(*i $Id: reductionops.mli 8708 2006-04-14 08:13:02Z jforest $ i*)
(*i*)
open Names
@@ -63,6 +63,7 @@ val nf_betaiota : local_reduction_function
val nf_betadeltaiota : reduction_function
val nf_evar : evar_map -> constr -> constr
+val nf_betaiotaevar_preserving_vm_cast : reduction_function
(* Lazy strategy, weak head reduction *)
val whd_evar : evar_map -> constr -> constr
val whd_beta : local_reduction_function
@@ -111,6 +112,10 @@ val whd_betadeltaiotaeta_stack : stack_reduction_function
val whd_betadeltaiotaeta_state : state_reduction_function
val whd_betadeltaiotaeta : reduction_function
+
+
+
+
val beta_applist : constr * constr list -> constr
val hnf_prod_app : env -> evar_map -> constr -> constr -> constr
@@ -121,12 +126,14 @@ val hnf_lam_appvect : env -> evar_map -> constr -> constr array -> constr
val hnf_lam_applist : env -> evar_map -> constr -> constr list -> constr
val splay_prod : env -> evar_map -> constr -> (name * constr) list * constr
+val splay_lambda : env -> evar_map -> constr -> (name * constr) list * constr
val splay_arity : env -> evar_map -> constr -> (name * constr) list * sorts
val sort_of_arity : env -> constr -> sorts
val decomp_n_prod :
env -> evar_map -> int -> constr -> Sign.rel_context * constr
val splay_prod_assum :
env -> evar_map -> constr -> Sign.rel_context * constr
+val decomp_sort : env -> evar_map -> types -> sorts
type 'a miota_args = {
mP : constr; (* the result type *)
@@ -162,10 +169,6 @@ val reduce_fix : local_state_reduction_function -> fixpoint
type conversion_test = constraints -> constraints
-type conv_pb =
- | CONV
- | CUMUL
-
val pb_is_equal : conv_pb -> bool
val pb_equal : conv_pb -> conv_pb
@@ -188,3 +191,7 @@ val instance : (metavariable * constr) list -> constr -> constr
val hnf : env -> 'a evar_map -> constr -> constr * constr list
i*)
val apprec : state_reduction_function
+
+(*s Meta-related reduction functions *)
+val meta_instance : evar_defs -> constr freelisted -> constr
+val nf_meta : evar_defs -> constr -> constr
diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml
index 061382f7..32da4cea 100644
--- a/pretyping/retyping.ml
+++ b/pretyping/retyping.ml
@@ -6,22 +6,17 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: retyping.ml,v 1.43.2.1 2004/07/16 19:30:46 herbelin Exp $ *)
+(* $Id: retyping.ml 8673 2006-03-29 21:21:52Z herbelin $ *)
open Util
open Term
open Inductive
+open Inductiveops
open Names
open Reductionops
open Environ
open Typeops
open Declarations
-open Instantiate
-
-let outsort env sigma t =
- match kind_of_term (whd_betadeltaiota env sigma t) with
- | Sort s -> s
- | _ -> anomaly "Retyping: found a type of type which is not a sort"
let rec subst_type env sigma typ = function
| [] -> typ
@@ -38,9 +33,9 @@ let rec subst_type env sigma typ = function
let sort_of_atomic_type env sigma ft args =
let rec concl_of_arity env ar =
match kind_of_term (whd_betadeltaiota env sigma ar) with
- | Prod (na, t, b) -> concl_of_arity (push_rel (na,None,t) env) b
- | Sort s -> s
- | _ -> outsort env sigma (subst_type env sigma ft (Array.to_list args))
+ | Prod (na, t, b) -> concl_of_arity (push_rel (na,None,t) env) b
+ | Sort s -> s
+ | _ -> decomp_sort env sigma (subst_type env sigma ft (Array.to_list args))
in concl_of_arity env ft
let typeur sigma metamap =
@@ -61,7 +56,7 @@ let typeur sigma metamap =
| Const c ->
let cb = lookup_constant c env in
body_of_type cb.const_type
- | Evar ev -> existential_type sigma ev
+ | Evar ev -> Evd.existential_type sigma ev
| Ind ind -> body_of_type (type_of_inductive env ind)
| Construct cstr -> body_of_type (type_of_constructor env cstr)
| Case (_,p,c,lf) ->
@@ -78,15 +73,18 @@ let typeur sigma metamap =
subst1 b (type_of (push_rel (name,Some b,c1) env) c2)
| Fix ((_,i),(_,tys,_)) -> tys.(i)
| CoFix (i,(_,tys,_)) -> tys.(i)
- | App(f,args)->
+ | App(f,args) when isInd f ->
+ let t = type_of_applied_inductive env (destInd f) args in
+ strip_outer_cast (subst_type env sigma t (Array.to_list args))
+ | App(f,args) ->
strip_outer_cast
(subst_type env sigma (type_of env f) (Array.to_list args))
- | Cast (c,t) -> t
+ | Cast (c,_, t) -> t
| Sort _ | Prod _ -> mkSort (sort_of env cstr)
and sort_of env t =
match kind_of_term t with
- | Cast (c,s) when isSort s -> destSort s
+ | Cast (c,_, s) when isSort s -> destSort s
| Sort (Prop c) -> type_0
| Sort (Type u) -> Type (Univ.super u)
| Prod (name,t,c2) ->
@@ -95,16 +93,21 @@ let typeur sigma metamap =
| Prop _, (Prop Pos as s) -> s
| Type _, (Prop Pos as s) when
Environ.engagement env = Some ImpredicativeSet -> s
- | Type _ as s, Prop Pos -> s
- | _, (Type u2 as s) -> s (*Type Univ.dummy_univ*))
+ | Type u1, Prop Pos -> Type (Univ.sup u1 Univ.base_univ)
+ | Prop Pos, (Type u2) -> Type (Univ.sup Univ.base_univ u2)
+ | Prop Null, (Type _ as s) -> s
+ | Type u1, Type u2 -> Type (Univ.sup u1 u2))
+ | App(f,args) when isInd f ->
+ let t = type_of_applied_inductive env (destInd f) args in
+ sort_of_atomic_type env sigma t args
| App(f,args) -> sort_of_atomic_type env sigma (type_of env f) args
| Lambda _ | Fix _ | Construct _ ->
anomaly "sort_of: Not a type (1)"
- | _ -> outsort env sigma (type_of env t)
+ | _ -> decomp_sort env sigma (type_of env t)
and sort_family_of env t =
match kind_of_term t with
- | Cast (c,s) when isSort s -> family_of_sort (destSort s)
+ | Cast (c,_, s) when isSort s -> family_of_sort (destSort s)
| Sort (Prop c) -> InType
| Sort (Type u) -> InType
| Prod (name,t,c2) -> sort_family_of (push_rel (name,None,t) env) c2
@@ -112,16 +115,31 @@ let typeur sigma metamap =
family_of_sort (sort_of_atomic_type env sigma (type_of env f) args)
| Lambda _ | Fix _ | Construct _ ->
anomaly "sort_of: Not a type (1)"
- | _ -> family_of_sort (outsort env sigma (type_of env t))
+ | _ -> family_of_sort (decomp_sort env sigma (type_of env t))
+
+ and type_of_applied_inductive env ind args =
+ let specif = lookup_mind_specif env ind in
+ let t = Inductive.type_of_inductive specif in
+ if is_small_inductive specif then
+ (* No polymorphism *)
+ t
+ else
+ (* Retyping constructor with the actual arguments *)
+ let env',llc,ls0 = constructor_instances env specif ind args in
+ let lls = Array.map (Array.map (sort_of env')) llc in
+ let ls = Array.map max_inductive_sort lls in
+ set_inductive_level env (find_inductive_level env specif ind ls0 ls) t
- in type_of, sort_of, sort_family_of
+ in type_of, sort_of, sort_family_of, type_of_applied_inductive
-let get_type_of env sigma c = let f,_,_ = typeur sigma [] in f env c
-let get_sort_of env sigma t = let _,f,_ = typeur sigma [] in f env t
-let get_sort_family_of env sigma c = let _,_,f = typeur sigma [] in f env c
+let get_type_of env sigma c = let f,_,_,_ = typeur sigma [] in f env c
+let get_sort_of env sigma t = let _,f,_,_ = typeur sigma [] in f env t
+let get_sort_family_of env sigma c = let _,_,f,_ = typeur sigma [] in f env c
+let type_of_applied_inductive env sigma ind args =
+ let _,_,_,f = typeur sigma [] in f env ind args
let get_type_of_with_meta env sigma metamap =
- let f,_,_ = typeur sigma metamap in f env
+ let f,_,_,_ = typeur sigma metamap in f env
(* Makes an assumption from a constr *)
let get_assumption_of env evc c = c
diff --git a/pretyping/retyping.mli b/pretyping/retyping.mli
index f29ac8d8..7adec66b 100644
--- a/pretyping/retyping.mli
+++ b/pretyping/retyping.mli
@@ -6,15 +6,13 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: retyping.mli,v 1.16.2.1 2004/07/16 19:30:46 herbelin Exp $ i*)
+(*i $Id: retyping.mli 8673 2006-03-29 21:21:52Z herbelin $ i*)
(*i*)
open Names
open Term
open Evd
open Environ
-open Pattern
-open Termops
(*i*)
(* This family of functions assumes its constr argument is known to be
@@ -23,14 +21,18 @@ open Termops
either produces a wrong result or raise an anomaly. Use with care.
It doesn't handle predicative universes too. *)
-val get_type_of : env -> evar_map -> constr -> constr
+val get_type_of : env -> evar_map -> constr -> types
val get_sort_of : env -> evar_map -> types -> sorts
val get_sort_family_of : env -> evar_map -> types -> sorts_family
-val get_type_of_with_meta : env -> evar_map -> metamap -> constr -> constr
+val get_type_of_with_meta :
+ env -> evar_map -> Termops.metamap -> constr -> types
(* Makes an assumption from a constr *)
val get_assumption_of : env -> evar_map -> constr -> types
(* Makes an unsafe judgment from a constr *)
val get_judgment_of : env -> evar_map -> constr -> unsafe_judgment
+
+val type_of_applied_inductive : env -> evar_map -> inductive ->
+ constr array -> types
diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml
index e8bde1f3..88af6290 100644
--- a/pretyping/tacred.ml
+++ b/pretyping/tacred.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: tacred.ml,v 1.75.2.7 2005/11/02 13:18:43 herbelin Exp $ *)
+(* $Id: tacred.ml 8003 2006-02-07 22:11:50Z herbelin $ *)
open Pp
open Util
@@ -20,39 +20,26 @@ open Inductive
open Environ
open Reductionops
open Closure
-open Instantiate
open Cbv
open Rawterm
-exception Elimconst
-exception Redelimination
+(* Errors *)
-let set_opaque_const = Conv_oracle.set_opaque_const
-let set_transparent_const sp =
- let cb = Global.lookup_constant sp in
- if cb.const_body <> None & cb.const_opaque then
- errorlabstrm "set_transparent_const"
- (str "Cannot make" ++ spc () ++
- Nametab.pr_global_env Idset.empty (ConstRef sp) ++
- spc () ++ str "transparent because it was declared opaque.");
- Conv_oracle.set_transparent_const sp
+type reduction_tactic_error =
+ InvalidAbstraction of env * constr * (env * Type_errors.type_error)
-let set_opaque_var = Conv_oracle.set_opaque_var
-let set_transparent_var = Conv_oracle.set_transparent_var
+exception ReductionTacticError of reduction_tactic_error
-let _ =
- Summary.declare_summary "Transparent constants and variables"
- { Summary.freeze_function = Conv_oracle.freeze;
- Summary.unfreeze_function = Conv_oracle.unfreeze;
- Summary.init_function = Conv_oracle.init;
- Summary.survive_module = false;
- Summary.survive_section = false }
+(* Evaluable reference *)
+
+exception Elimconst
+exception Redelimination
let is_evaluable env ref =
match ref with
EvalConstRef kn ->
let (ids,kns) = Conv_oracle.freeze() in
- KNpred.mem kn kns &
+ Cpred.mem kn kns &
let cb = Environ.lookup_constant kn env in
cb.const_body <> None & not cb.const_opaque
| EvalVarRef id ->
@@ -84,7 +71,7 @@ let destEvalRef c = match kind_of_term c with
| Var id -> EvalVar id
| Rel n -> EvalRel n
| Evar ev -> EvalEvar ev
- | _ -> anomaly "Not an evaluable reference"
+ | _ -> anomaly "Not an unfoldable reference"
let reference_opt_value sigma env = function
| EvalConst cst -> constant_opt_value env cst
@@ -94,7 +81,7 @@ let reference_opt_value sigma env = function
| EvalRel n ->
let (_,v,_) = lookup_rel n env in
option_app (lift n) v
- | EvalEvar ev -> existential_opt_value sigma ev
+ | EvalEvar ev -> Evd.existential_opt_value sigma ev
exception NotEvaluable
let reference_value sigma env c =
@@ -145,17 +132,21 @@ let _ =
Summary.survive_module = false;
Summary.survive_section = false }
-
(* Check that c is an "elimination constant"
- [xn:An]..[x1:A1](<P>MutCase (Rel i) of f1..fk end g1 ..gp)
- or [xn:An]..[x1:A1](Fix(f|t) (Rel i1) ..(Rel ip))
- with i1..ip distinct variables not occuring in t
- keep relevenant information ([i1,Ai1;..;ip,Aip],n,b)
- with b = true in case of a fixpoint in order to compute
- an equivalent of Fix(f|t)[xi<-ai] as
- [yip:Bip]..[yi1:Bi1](F bn..b1)
- == [yip:Bip]..[yi1:Bi1](Fix(f|t)[xi<-ai] (Rel 1)..(Rel p))
- with bj=aj if j<>ik and bj=(Rel c) and Bic=Aic[xn..xic-1 <- an..aic-1] *)
+
+ either [xn:An]..[x1:A1](<P>Case (Rel i) of f1..fk end g1 ..gp)
+
+ or [xn:An]..[x1:A1](Fix(f|t) (Rel i1) ..(Rel ip))
+ with i1..ip distinct variables not occuring in t
+
+ In the second case, keep relevenant information ([i1,Ai1;..;ip,Aip],n)
+ in order to compute an equivalent of Fix(f|t)[xi<-ai] as
+
+ [yip:Bip]..[yi1:Bi1](F bn..b1)
+ == [yip:Bip]..[yi1:Bi1](Fix(f|t)[xi<-ai] (Rel p)..(Rel 1))
+
+ with bj=aj if j<>ik and bj=(Rel c) and Bic=Aic[xn..xic-1 <- an..aic-1]
+*)
let check_fix_reversibility labs args ((lv,i),(_,tys,bds)) =
let n = List.length labs in
@@ -199,8 +190,8 @@ let invert_name labs l na0 env sigma ref = function
| EvalRel _ | EvalEvar _ -> None
| EvalVar id' -> Some (EvalVar id)
| EvalConst kn ->
- let (mp,dp,_) = repr_kn kn in
- Some (EvalConst (make_kn mp dp (label_of_id id))) in
+ let (mp,dp,_) = repr_con kn in
+ Some (EvalConst (make_con mp dp (label_of_id id))) in
match refi with
| None -> None
| Some ref ->
@@ -242,7 +233,7 @@ let compute_consteval_mutual_fix sigma env ref =
match kind_of_term c' with
| Lambda (na,t,g) when l=[] ->
srec (push_rel (na,None,t) env) (minarg+1) (t::labs) ref g
- | Fix ((lv,i),(names,_,_) as fix) ->
+ | Fix ((lv,i),(names,_,_)) ->
(* Last known constant wrapping Fix is ref = [labs](Fix l) *)
(match compute_consteval_direct sigma env ref with
| NotAnElimination -> (*Above const was eliminable but this not!*)
@@ -294,43 +285,49 @@ let rev_firstn_liftn fn ln =
in
rfprec fn []
-(* EliminationFix ([(yi1,Ti1);...;(yip,Tip)],n) means f is some
- [y1:T1,...,yn:Tn](Fix(..) yi1 ... yip);
- f is applied to largs and we need for recursive calls to build
- [x1:Ti1',...,xp:Tip'](f a1..a(n-p) yi1 ... yip)
- where a1...an are the n first arguments of largs and Tik' is Tik[yil=al]
- To check ... *)
+(* If f is bound to EliminationFix (n',infos), then n' is the minimal
+ number of args for starting the reduction and infos is
+ (nbfix,[(yi1,Ti1);...;(yip,Tip)],n) indicating that f converts
+ to some [y1:T1,...,yn:Tn](Fix(..) yip .. yi1) where we can remark that
+ yij = Rel(n+1-j)
+
+ f is applied to largs and we need for recursive calls to build the function
+ g := [xp:Tip',...,x1:Ti1'](f a1 ... an)
+
+ s.t. (g u1 ... up) reduces to (Fix(..) u1 ... up)
+
+ This is made possible by setting
+ a_k:=z_j if k=i_j
+ a_k:=y_k otherwise
+
+ The type Tij' is Tij[yn..yi(j-1)..y1 <- ai(j-1)..a1]
+*)
+
+let x = Name (id_of_string "x")
let make_elim_fun (names,(nbfix,lv,n)) largs =
- let labs = list_firstn n (list_of_stack largs) in
+ let lu = list_firstn n (list_of_stack largs) in
let p = List.length lv in
- let ylv = List.map fst lv in
- let la' = list_map_i
- (fun q aq ->
- try (mkRel (p+1-(list_index (n-q) ylv)))
- with Not_found -> aq) 0
- (List.map (lift p) labs)
+ let lyi = List.map fst lv in
+ let la =
+ list_map_i (fun q aq ->
+ (* k from the comment is q+1 *)
+ try mkRel (p+1-(list_index (n-q) lyi))
+ with Not_found -> aq)
+ 0 (List.map (lift p) lu)
in
fun i ->
match names.(i) with
| None -> None
- | Some ref -> Some (
-(* let fi =
- if nbfix = 1 then
- mkEvalRef ref
- else
- match ref with
- | EvalConst (sp,args) ->
- mkConst (make_path (dirpath sp) id (kind_of_path sp),args)
- | _ -> anomaly "elimination of local fixpoints not implemented"
- in
-*)
- list_fold_left_i
- (fun i c (k,a) ->
- mkLambda (Name(id_of_string"x"),
- substl (rev_firstn_liftn (n-k) (-i) la') a,
- c))
- 1 (applistc (mkEvalRef ref) la') (List.rev lv))
+ | Some ref ->
+ let body = applistc (mkEvalRef ref) la in
+ let g =
+ list_fold_left_i (fun q (* j from comment is n+1-q *) c (ij,tij) ->
+ let subst = List.map (lift (-q)) (list_firstn (n-ij) la) in
+ let tij' = substl (List.rev subst) tij in
+ mkLambda (x,tij',c)
+ ) 1 body (List.rev lv)
+ in Some g
(* [f] is convertible to [Fix(recindices,bodynum),bodyvect)] make
the reduction using this extra information *)
@@ -372,7 +369,7 @@ let contract_cofix_use_function f (bodynum,(_,names,bodies as typedbodies)) =
let reduce_mind_case_use_function func env mia =
match kind_of_term mia.mconstr with
- | Construct(ind_sp,i as cstr_sp) ->
+ | Construct(ind_sp,i) ->
let real_cargs = list_skipn mia.mci.ci_npar mia.mcargs in
applist (mia.mlf.(i-1), real_cargs)
| CoFix (_,(names,_,_) as cofix) ->
@@ -380,8 +377,8 @@ let reduce_mind_case_use_function func env mia =
match names.(i) with
| Name id ->
if isConst func then
- let (mp,dp,_) = repr_kn (destConst func) in
- let kn = make_kn mp dp (label_of_id id) in
+ let (mp,dp,_) = repr_con (destConst func) in
+ let kn = make_con mp dp (label_of_id id) in
(match constant_opt_value env kn with
| None -> None
| Some _ -> Some (mkConst kn))
@@ -452,7 +449,7 @@ let rec red_elim_const env sigma ref largs =
and construct_const env sigma =
let rec hnfstack (x, stack as s) =
match kind_of_term x with
- | Cast (c,_) -> hnfstack (c, stack)
+ | Cast (c,_,_) -> hnfstack (c, stack)
| App (f,cl) -> hnfstack (f, append_stack cl stack)
| Lambda (id,t,c) ->
(match decomp_stack stack with
@@ -491,7 +488,7 @@ and construct_const env sigma =
(* Red reduction tactic: reduction to a product *)
-let internal_red_product env sigma c =
+let try_red_product env sigma c =
let simpfun = clos_norm_flags betaiotazeta env sigma in
let rec redrec env x =
match kind_of_term x with
@@ -506,7 +503,7 @@ let internal_red_product env sigma c =
let stack' = stack_assign stack recargnum recarg' in
simpfun (app_stack (f,stack')))
| _ -> simpfun (appvect (redrec env f, l)))
- | Cast (c,_) -> redrec env c
+ | Cast (c,_,_) -> redrec env c
| Prod (x,a,b) -> mkProd (x, a, redrec (push_rel (x,None,a) env) b)
| LetIn (x,a,b,t) -> redrec env (subst1 a t)
| Case (ci,p,d,lf) -> simpfun (mkCase (ci,p,redrec env d,lf))
@@ -521,11 +518,9 @@ let internal_red_product env sigma c =
in redrec env c
let red_product env sigma c =
- try internal_red_product env sigma c
+ try try_red_product env sigma c
with Redelimination -> error "Not reducible"
-(* Hnf reduction tactic: *)
-
let hnf_constr env sigma c =
let rec redrec (x, largs as s) =
match kind_of_term x with
@@ -536,7 +531,7 @@ let hnf_constr env sigma c =
stacklam redrec [a] c rest)
| LetIn (n,b,t,c) -> stacklam redrec [b] c largs
| App (f,cl) -> redrec (f, append_stack cl largs)
- | Cast (c,_) -> redrec (c, largs)
+ | Cast (c,_,_) -> redrec (c, largs)
| Case (ci,p,c,lf) ->
(try
redrec
@@ -577,7 +572,7 @@ let whd_nf env sigma c =
stacklam nf_app [a1] c2 rest)
| LetIn (n,b,t,c) -> stacklam nf_app [b] c stack
| App (f,cl) -> nf_app (f, append_stack cl stack)
- | Cast (c,_) -> nf_app (c, stack)
+ | Cast (c,_,_) -> nf_app (c, stack)
| Case (ci,p,d,lf) ->
(try
nf_app (special_red_case sigma env nf_app (ci,p,d,lf), stack)
@@ -598,9 +593,6 @@ let whd_nf env sigma c =
let nf env sigma c = strong whd_nf env sigma c
-let is_reference c =
- try let r = reference_of_constr c in true with _ -> false
-
let is_head c t =
match kind_of_term t with
| App (f,_) -> f = c
@@ -609,7 +601,6 @@ let is_head c t =
let contextually byhead (locs,c) f env sigma t =
let maxocc = List.fold_right max locs 0 in
let pos = ref 1 in
- let check = ref true in
let except = List.exists (fun n -> n<0) locs in
if except & (List.exists (fun n -> n>=0) locs)
then error "mixing of positive and negative occurences"
@@ -626,7 +617,7 @@ let contextually byhead (locs,c) f env sigma t =
f env sigma t
else if byhead then
(* find other occurrences of c in t; TODO: ensure left-to-right *)
- let (f,l) = destApplication t in
+ let (f,l) = destApp t in
mkApp (f, array_map_left (traverse envc) l)
else
t
@@ -637,7 +628,7 @@ let contextually byhead (locs,c) f env sigma t =
in
let t' = traverse (env,c) t in
if locs <> [] & List.exists (fun o -> o >= !pos or o <= - !pos) locs then
- errorlabstrm "contextually" (str "Too few occurences");
+ error_invalid_occurrence locs;
t'
(* linear bindings (following pretty-printer) of the value of name in c.
@@ -652,7 +643,7 @@ let rec substlin env name n ol c =
with
NotEvaluableConst _ ->
errorlabstrm "substlin"
- (pr_kn kn ++ str " is not a defined constant")
+ (pr_con kn ++ str " is not a defined constant")
else
((n+1), ol, c)
@@ -701,7 +692,7 @@ let rec substlin env name n ol c =
let (n2,ol2,c2') = substlin env name n1 ol1 c2 in
(n2,ol2,mkProd (na,c1',c2')))
- | Case (ci,p,d,llf) ->
+ | Case (ci,p,d,llf) ->
let rec substlist nn oll = function
| [] -> (nn,oll,[])
| f::lfe ->
@@ -712,24 +703,25 @@ let rec substlin env name n ol c =
let (nn2,oll2,lfe') = substlist nn1 oll1 lfe in
(nn2,oll2,f'::lfe'))
in
- let (n1,ol1,p') = substlin env name n ol p in (* ATTENTION ERREUR *)
- (match ol1 with (* si P pas affiche *)
- | [] -> (n1,[],mkCase (ci, p', d, llf))
+ (* p is printed after d in v8 syntax *)
+ let (n1,ol1,d') = substlin env name n ol d in
+ (match ol1 with
+ | [] -> (n1,[],mkCase (ci, p, d', llf))
| _ ->
- let (n2,ol2,d') = substlin env name n1 ol1 d in
+ let (n2,ol2,p') = substlin env name n1 ol1 p in
(match ol2 with
| [] -> (n2,[],mkCase (ci, p', d', llf))
| _ ->
let (n3,ol3,lf') = substlist n2 ol2 (Array.to_list llf)
in (n3,ol3,mkCase (ci, p', d', Array.of_list lf'))))
- | Cast (c1,c2) ->
+ | Cast (c1,k,c2) ->
let (n1,ol1,c1') = substlin env name n ol c1 in
(match ol1 with
- | [] -> (n1,[],mkCast (c1',c2))
- | _ ->
- let (n2,ol2,c2') = substlin env name n1 ol1 c2 in
- (n2,ol2,mkCast (c1',c2')))
+ | [] -> (n1,[],mkCast (c1',k,c2))
+ | _ ->
+ let (n2,ol2,c2') = substlin env name n1 ol1 c2 in
+ (n2,ol2,mkCast (c1',k,c2')))
| Fix _ ->
(warning "do not consider occurrences inside fixpoints"; (n,ol,c))
@@ -764,8 +756,8 @@ let unfoldoccs env sigma (occl,name) c =
| (_,[],uc) -> nf_betaiota uc
| (1,_,_) ->
error ((string_of_evaluable_ref env name)^" does not occur")
- | _ -> error ("bad occurrence numbers of "
- ^(string_of_evaluable_ref env name))
+ | (_,l,_) ->
+ error_invalid_occurrence l
(* Unfold reduction tactic: *)
let unfoldn loccname env sigma c =
@@ -797,78 +789,29 @@ let compute = cbv_betadeltaiota
(* gives [na:ta]c' such that c converts to ([na:ta]c' a), abstracting only
* the specified occurrences. *)
-let abstract_scheme env sigma (locc,a) t =
+let abstract_scheme env sigma (locc,a) c =
let ta = Retyping.get_type_of env sigma a in
let na = named_hd env ta Anonymous in
if occur_meta ta then error "cannot find a type for the generalisation";
if occur_meta a then
- mkLambda (na,ta,t)
+ mkLambda (na,ta,c)
else
- mkLambda (na, ta,subst_term_occ locc a t)
-
+ mkLambda (na,ta,subst_term_occ locc a c)
let pattern_occs loccs_trm env sigma c =
let abstr_trm = List.fold_right (abstract_scheme env sigma) loccs_trm c in
- let _ = Typing.type_of env sigma abstr_trm in
- applist(abstr_trm, List.map snd loccs_trm)
-
-(* Generic reduction: reduction functions used in reduction tactics *)
-
-type red_expr = (constr, evaluable_global_reference) red_expr_gen
-
-open RedFlags
-
-let make_flag_constant = function
- | EvalVarRef id -> fVAR id
- | EvalConstRef sp -> fCONST sp
-
-let make_flag f =
- let red = no_red in
- let red = if f.rBeta then red_add red fBETA else red in
- let red = if f.rIota then red_add red fIOTA else red in
- let red = if f.rZeta then red_add red fZETA else red in
- let red =
- if f.rDelta then (* All but rConst *)
- let red = red_add red fDELTA in
- let red = red_add_transparent red (Conv_oracle.freeze ()) in
- List.fold_right
- (fun v red -> red_sub red (make_flag_constant v))
- f.rConst red
- else (* Only rConst *)
- let red = red_add_transparent (red_add red fDELTA) all_opaque in
- List.fold_right
- (fun v red -> red_add red (make_flag_constant v))
- f.rConst red
- in red
-
-let red_expr_tab = ref Stringmap.empty
-
-let declare_red_expr s f =
- try
- let _ = Stringmap.find s !red_expr_tab in
- error ("There is already a reduction expression of name "^s)
- with Not_found ->
- red_expr_tab := Stringmap.add s f !red_expr_tab
-
-let reduction_of_redexp = function
- | Red internal -> if internal then internal_red_product else red_product
- | Hnf -> hnf_constr
- | Simpl (Some (_,c as lp)) -> contextually (is_reference c) lp nf
- | Simpl None -> nf
- | Cbv f -> cbv_norm_flags (make_flag f)
- | Lazy f -> clos_norm_flags (make_flag f)
- | Unfold ubinds -> unfoldn ubinds
- | Fold cl -> fold_commands cl
- | Pattern lp -> pattern_occs lp
- | ExtraRedExpr s ->
- (try Stringmap.find s !red_expr_tab
- with Not_found -> error("unknown user-defined reduction \""^s^"\""))
+ try
+ let _ = Typing.type_of env sigma abstr_trm in
+ applist(abstr_trm, List.map snd loccs_trm)
+ with Type_errors.TypeError (env',t) ->
+ raise (ReductionTacticError (InvalidAbstraction (env,abstr_trm,(env',t))))
+
(* Used in several tactics. *)
exception NotStepReducible
let one_step_reduce env sigma c =
- let rec redrec (x, largs as s) =
+ let rec redrec (x, largs) =
match kind_of_term x with
| Lambda (n,t,c) ->
(match decomp_stack largs with
@@ -885,7 +828,7 @@ let one_step_reduce env sigma c =
(match reduce_fix (whd_betadeltaiota_state env sigma) fix largs with
| Reduced s' -> s'
| NotReducible -> raise NotStepReducible)
- | Cast (c,_) -> redrec (c,largs)
+ | Cast (c,_,_) -> redrec (c,largs)
| _ when isEvalRef env x ->
let ref =
try destEvalRef x
@@ -940,14 +883,14 @@ let reduce_to_ref_gen allow_product env sigma ref t =
(str"Not an induction object of atomic type")
| _ ->
try
- if reference_of_constr c = ref
+ if global_of_constr c = ref
then it_mkProd_or_LetIn t l
else raise Not_found
with Not_found ->
try
let t' = nf_betaiota (one_step_reduce env sigma t) in
elimrec env t' l
- with NotStepReducible ->
+ with NotStepReducible ->
errorlabstrm ""
(str "Not a statement of conclusion " ++
Nametab.pr_global_env Idset.empty ref)
diff --git a/pretyping/tacred.mli b/pretyping/tacred.mli
index 7998a8fd..a5468435 100644
--- a/pretyping/tacred.mli
+++ b/pretyping/tacred.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: tacred.mli,v 1.21.2.2 2005/01/21 16:42:37 herbelin Exp $ i*)
+(*i $Id: tacred.mli 8003 2006-02-07 22:11:50Z herbelin $ i*)
(*i*)
open Names
@@ -18,15 +18,23 @@ open Closure
open Rawterm
(*i*)
+type reduction_tactic_error =
+ InvalidAbstraction of env * constr * (env * Type_errors.type_error)
+
+exception ReductionTacticError of reduction_tactic_error
+
(*s Reduction functions associated to tactics. \label{tacred} *)
val is_evaluable : env -> evaluable_global_reference -> bool
exception Redelimination
-(* Red (raise Redelimination if nothing reducible) *)
+(* Red (raise user error if nothing reducible) *)
val red_product : reduction_function
+(* Red (raise Redelimination if nothing reducible) *)
+val try_red_product : reduction_function
+
(* Hnf *)
val hnf_constr : reduction_function
@@ -69,17 +77,5 @@ val reduce_to_quantified_ref :
val reduce_to_atomic_ref :
env -> evar_map -> Libnames.global_reference -> types -> types
-type red_expr = (constr, evaluable_global_reference) red_expr_gen
-
val contextually : bool -> constr occurrences -> reduction_function
-> reduction_function
-val reduction_of_redexp : red_expr -> reduction_function
-
-val declare_red_expr : string -> reduction_function -> unit
-
-(* Opaque and Transparent commands. *)
-val set_opaque_const : constant -> unit
-val set_transparent_const : constant -> unit
-
-val set_opaque_var : identifier -> unit
-val set_transparent_var : identifier -> unit
diff --git a/pretyping/termops.ml b/pretyping/termops.ml
index 8f12ca62..89de5537 100644
--- a/pretyping/termops.ml
+++ b/pretyping/termops.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: termops.ml,v 1.29.2.1 2004/07/16 19:30:46 herbelin Exp $ *)
+(* $Id: termops.ml 8003 2006-02-07 22:11:50Z herbelin $ *)
open Pp
open Util
@@ -35,60 +35,116 @@ let pr_name = function
| Anonymous -> str "_"
let pr_sp sp = str(string_of_kn sp)
+let pr_con sp = str(string_of_con sp)
-let rec print_constr c = match kind_of_term c with
+let rec pr_constr c = match kind_of_term c with
| Rel n -> str "#"++int n
| Meta n -> str "Meta(" ++ int n ++ str ")"
| Var id -> pr_id id
| Sort s -> print_sort s
- | Cast (c,t) -> hov 1
- (str"(" ++ print_constr c ++ cut() ++
- str":" ++ print_constr t ++ str")")
+ | Cast (c,_, t) -> hov 1
+ (str"(" ++ pr_constr c ++ cut() ++
+ str":" ++ pr_constr t ++ str")")
| Prod (Name(id),t,c) -> hov 1
- (str"forall " ++ pr_id id ++ str":" ++ print_constr t ++ str"," ++
- spc() ++ print_constr c)
+ (str"forall " ++ pr_id id ++ str":" ++ pr_constr t ++ str"," ++
+ spc() ++ pr_constr c)
| Prod (Anonymous,t,c) -> hov 0
- (str"(" ++ print_constr t ++ str " ->" ++ spc() ++
- print_constr c ++ str")")
+ (str"(" ++ pr_constr t ++ str " ->" ++ spc() ++
+ pr_constr c ++ str")")
| Lambda (na,t,c) -> hov 1
(str"fun " ++ pr_name na ++ str":" ++
- print_constr t ++ str" =>" ++ spc() ++ print_constr c)
+ pr_constr t ++ str" =>" ++ spc() ++ pr_constr c)
| LetIn (na,b,t,c) -> hov 0
- (str"let " ++ pr_name na ++ str":=" ++ print_constr b ++
- str":" ++ brk(1,2) ++ print_constr t ++ cut() ++
- print_constr c)
+ (str"let " ++ pr_name na ++ str":=" ++ pr_constr b ++
+ str":" ++ brk(1,2) ++ pr_constr t ++ cut() ++
+ pr_constr c)
| App (c,l) -> hov 1
- (str"(" ++ print_constr c ++ spc() ++
- prlist_with_sep spc print_constr (Array.to_list l) ++ str")")
+ (str"(" ++ pr_constr c ++ spc() ++
+ prlist_with_sep spc pr_constr (Array.to_list l) ++ str")")
| Evar (e,l) -> hov 1
(str"Evar#" ++ int e ++ str"{" ++
- prlist_with_sep spc print_constr (Array.to_list l) ++str"}")
- | Const c -> str"Cst(" ++ pr_sp c ++ str")"
+ prlist_with_sep spc pr_constr (Array.to_list l) ++str"}")
+ | Const c -> str"Cst(" ++ pr_con c ++ str")"
| Ind (sp,i) -> str"Ind(" ++ pr_sp sp ++ str"," ++ int i ++ str")"
| Construct ((sp,i),j) ->
str"Constr(" ++ pr_sp sp ++ str"," ++ int i ++ str"," ++ int j ++ str")"
| Case (ci,p,c,bl) -> v 0
- (hv 0 (str"<"++print_constr p++str">"++ cut() ++ str"Case " ++
- print_constr c ++ str"of") ++ cut() ++
- prlist_with_sep (fun _ -> brk(1,2)) print_constr (Array.to_list bl) ++
+ (hv 0 (str"<"++pr_constr p++str">"++ cut() ++ str"Case " ++
+ pr_constr c ++ str"of") ++ cut() ++
+ prlist_with_sep (fun _ -> brk(1,2)) pr_constr (Array.to_list bl) ++
cut() ++ str"end")
| Fix ((t,i),(lna,tl,bl)) ->
let fixl = Array.mapi (fun i na -> (na,t.(i),tl.(i),bl.(i))) lna in
hov 1
(str"fix " ++ int i ++ spc() ++ str"{" ++
v 0 (prlist_with_sep spc (fun (na,i,ty,bd) ->
- pr_name na ++ str"/" ++ int i ++ str":" ++ print_constr ty ++
- cut() ++ str":=" ++ print_constr bd) (Array.to_list fixl)) ++
+ pr_name na ++ str"/" ++ int i ++ str":" ++ pr_constr ty ++
+ cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
str"}")
| CoFix(i,(lna,tl,bl)) ->
let fixl = Array.mapi (fun i na -> (na,tl.(i),bl.(i))) lna in
hov 1
(str"cofix " ++ int i ++ spc() ++ str"{" ++
v 0 (prlist_with_sep spc (fun (na,ty,bd) ->
- pr_name na ++ str":" ++ print_constr ty ++
- cut() ++ str":=" ++ print_constr bd) (Array.to_list fixl)) ++
+ pr_name na ++ str":" ++ pr_constr ty ++
+ cut() ++ str":=" ++ pr_constr bd) (Array.to_list fixl)) ++
str"}")
+let term_printer = ref (fun _ -> pr_constr)
+let print_constr_env t = !term_printer t
+let print_constr t = !term_printer (Global.env()) t
+let set_print_constr f = term_printer := f
+
+let pr_var_decl env (id,c,typ) =
+ let pbody = match c with
+ | None -> (mt ())
+ | Some c ->
+ (* Force evaluation *)
+ let pb = print_constr_env env c in
+ (str" := " ++ pb ++ cut () ) in
+ let pt = print_constr_env env typ in
+ let ptyp = (str" : " ++ pt) in
+ (pr_id id ++ hov 0 (pbody ++ ptyp))
+
+let pr_rel_decl env (na,c,typ) =
+ let pbody = match c with
+ | None -> mt ()
+ | Some c ->
+ (* Force evaluation *)
+ let pb = print_constr_env env c in
+ (str":=" ++ spc () ++ pb ++ spc ()) in
+ let ptyp = print_constr_env env typ in
+ match na with
+ | Anonymous -> hov 0 (str"<>" ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
+ | Name id -> hov 0 (pr_id id ++ spc () ++ pbody ++ str":" ++ spc () ++ ptyp)
+
+let print_named_context env =
+ hv 0 (fold_named_context
+ (fun env d pps ->
+ pps ++ ws 2 ++ pr_var_decl env d)
+ env ~init:(mt ()))
+
+let print_rel_context env =
+ hv 0 (fold_rel_context
+ (fun env d pps -> pps ++ ws 2 ++ pr_rel_decl env d)
+ env ~init:(mt ()))
+
+let print_env env =
+ let sign_env =
+ fold_named_context
+ (fun env d pps ->
+ let pidt = pr_var_decl env d in
+ (pps ++ fnl () ++ pidt))
+ env ~init:(mt ())
+ in
+ let db_env =
+ fold_rel_context
+ (fun env d pps ->
+ let pnat = pr_rel_decl env d in (pps ++ fnl () ++ pnat))
+ env ~init:(mt ())
+ in
+ (sign_env ++ db_env)
+
(*let current_module = ref empty_dirpath
let set_module m = current_module := m*)
@@ -98,6 +154,20 @@ let new_univ =
(fun sp ->
incr univ_gen;
Univ.make_univ (Lib.library_dp(),!univ_gen))
+let new_Type () = mkType (new_univ ())
+let new_Type_sort () = Type (new_univ ())
+
+(* This refreshes universes in types; works only for inferred types (i.e. for
+ types of the form (x1:A1)...(xn:An)B with B a sort or an atom in
+ head normal form) *)
+let refresh_universes t =
+ let modified = ref false in
+ let rec refresh t = match kind_of_term t with
+ | Sort (Type _) -> modified := true; new_Type ()
+ | Prod (na,u,v) -> mkProd (na,u,refresh v)
+ | _ -> t in
+ let t' = refresh t in
+ if !modified then t' else t
let new_sort_in_family = function
| InProp -> mk_Prop
@@ -185,6 +255,8 @@ let it_named_context_quantifier f ~init =
let it_mkNamedProd_or_LetIn = it_named_context_quantifier mkNamedProd_or_LetIn
let it_mkNamedLambda_or_LetIn = it_named_context_quantifier mkNamedLambda_or_LetIn
+let it_mkNamedProd_wo_LetIn = it_named_context_quantifier mkNamedProd_wo_LetIn
+
(* *)
(* strips head casts and flattens head applications *)
@@ -192,11 +264,11 @@ let rec strip_head_cast c = match kind_of_term c with
| App (f,cl) ->
let rec collapse_rec f cl2 = match kind_of_term f with
| App (g,cl1) -> collapse_rec g (Array.append cl1 cl2)
- | Cast (c,_) -> collapse_rec c cl2
- | _ -> if cl2 = [||] then f else mkApp (f,cl2)
+ | Cast (c,_,_) -> collapse_rec c cl2
+ | _ -> if Array.length cl2 = 0 then f else mkApp (f,cl2)
in
collapse_rec f cl
- | Cast (c,t) -> strip_head_cast c
+ | Cast (c,_,_) -> strip_head_cast c
| _ -> c
(* [map_constr_with_named_binders g f l c] maps [f l] on the immediate
@@ -208,7 +280,7 @@ let rec strip_head_cast c = match kind_of_term c with
let map_constr_with_named_binders g f l c = match kind_of_term c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> c
- | Cast (c,t) -> mkCast (f l c, f l t)
+ | Cast (c,k,t) -> mkCast (f l c, k, f l t)
| Prod (na,t,c) -> mkProd (na, f l t, f (g na l) c)
| Lambda (na,t,c) -> mkLambda (na, f l t, f (g na l) c)
| LetIn (na,b,t,c) -> mkLetIn (na, f l b, f l t, f (g na l) c)
@@ -243,7 +315,7 @@ let fold_rec_types g (lna,typarray,_) e =
let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> c
- | Cast (c,t) -> let c' = f l c in mkCast (c', f l t)
+ | Cast (c,k,t) -> let c' = f l c in mkCast (c',k,f l t)
| Prod (na,t,c) ->
let t' = f l t in
mkProd (na, t', f (g (na,None,t) l) c)
@@ -263,7 +335,9 @@ let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
mkApp (hd, [| f l a |])
| Evar (e,al) -> mkEvar (e, array_map_left (f l) al)
| Case (ci,p,c,bl) ->
- let p' = f l p in let c' = f l c in
+ (* In v8 concrete syntax, predicate is after the term to match! *)
+ let c' = f l c in
+ let p' = f l p in
mkCase (ci, p', c', array_map_left (f l) bl)
| Fix (ln,(lna,tl,bl as fx)) ->
let l' = fold_rec_types g fx l in
@@ -278,10 +352,10 @@ let map_constr_with_binders_left_to_right g f l c = match kind_of_term c with
let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> cstr
- | Cast (c,t) ->
+ | Cast (c,k, t) ->
let c' = f l c in
let t' = f l t in
- if c==c' && t==t' then cstr else mkCast (c', t')
+ if c==c' && t==t' then cstr else mkCast (c', k, t')
| Prod (na,t,c) ->
let t' = f l t in
let c' = f (g (na,None,t) l) c in
@@ -335,7 +409,7 @@ let map_constr_with_full_binders g f l cstr = match kind_of_term cstr with
let fold_constr_with_binders g f n acc c = match kind_of_term c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> acc
- | Cast (c,t) -> f n (f n acc c) t
+ | Cast (c,_, t) -> f n (f n acc c) t
| Prod (_,t,c) -> f (g n) (f n acc t) c
| Lambda (_,t,c) -> f (g n) (f n acc t) c
| LetIn (_,b,t,c) -> f (g n) (f n (f n acc b) t) c
@@ -359,7 +433,7 @@ let fold_constr_with_binders g f n acc c = match kind_of_term c with
let iter_constr_with_full_binders g f l c = match kind_of_term c with
| (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _
| Construct _) -> ()
- | Cast (c,t) -> f l c; f l t
+ | Cast (c,_, t) -> f l c; f l t
| Prod (na,t,c) -> f l t; f (g (na,None,t) l) c
| Lambda (na,t,c) -> f l t; f (g (na,None,t) l) c
| LetIn (na,b,t,c) -> f l b; f l t; f (g (na,Some b,t) l) c
@@ -541,10 +615,14 @@ let replace_term = replace_term_gen eq_constr
bindings is done. The list may contain only negative occurrences
that will not be substituted. *)
+let error_invalid_occurrence l =
+ errorlabstrm ""
+ (str ("Invalid occurrence " ^ plural (List.length l) "number" ^": ") ++
+ prlist_with_sep spc int l)
+
let subst_term_occ_gen locs occ c t =
let maxocc = List.fold_right max locs 0 in
let pos = ref occ in
- let check = ref true in
let except = List.exists (fun n -> n<0) locs in
if except & (List.exists (fun n -> n>=0) locs)
then error "mixing of positive and negative occurences"
@@ -573,8 +651,8 @@ let subst_term_occ locs c t =
t
else
let (nbocc,t') = subst_term_occ_gen locs 1 c t in
- if List.exists (fun o -> o >= nbocc or o <= -nbocc) locs then
- errorlabstrm "subst_term_occ" (str "Too few occurences");
+ let rest = List.filter (fun o -> o >= nbocc or o <= -nbocc) locs in
+ if rest <> [] then error_invalid_occurrence rest;
t'
let subst_term_occ_decl locs c (id,bodyopt,typ as d) =
@@ -588,8 +666,8 @@ let subst_term_occ_decl locs c (id,bodyopt,typ as d) =
else
let (nbocc,body') = subst_term_occ_gen locs 1 c body in
let (nbocc',t') = subst_term_occ_gen locs nbocc c typ in
- if List.exists (fun o -> o >= nbocc' or o <= -nbocc') locs then
- errorlabstrm "subst_term_occ_decl" (str "Too few occurences");
+ let rest = List.filter (fun o -> o >= nbocc' or o <= -nbocc') locs in
+ if rest <> [] then error_invalid_occurrence rest;
(id,Some body',t')
@@ -626,10 +704,10 @@ let hdchar env c =
| Prod (_,_,c) -> hdrec (k+1) c
| Lambda (_,_,c) -> hdrec (k+1) c
| LetIn (_,_,_,c) -> hdrec (k+1) c
- | Cast (c,_) -> hdrec k c
+ | Cast (c,_,_) -> hdrec k c
| App (f,l) -> hdrec k f
| Const kn ->
- let c = lowercase_first_char (id_of_label (label kn)) in
+ let c = lowercase_first_char (id_of_label (con_label kn)) in
if c = "?" then "y" else c
| Ind ((kn,i) as x) ->
if i=0 then
@@ -667,8 +745,11 @@ let named_hd env a = function
let named_hd_type env a = named_hd env (body_of_type a)
-let prod_name env (n,a,b) = mkProd (named_hd_type env a n, a, b)
-let lambda_name env (n,a,b) = mkLambda (named_hd_type env a n, a, b)
+let mkProd_name env (n,a,b) = mkProd (named_hd_type env a n, a, b)
+let mkLambda_name env (n,a,b) = mkLambda (named_hd_type env a n, a, b)
+
+let lambda_name = mkLambda_name
+let prod_name = mkProd_name
let prod_create env (a,b) = mkProd (named_hd_type env a Anonymous, a, b)
let lambda_create env (a,b) = mkLambda (named_hd_type env a Anonymous, a, b)
@@ -714,6 +795,7 @@ let ids_of_rel_context sign =
Sign.fold_rel_context
(fun (na,_,_) l -> match na with Name id -> id::l | Anonymous -> l)
sign ~init:[]
+
let ids_of_named_context sign =
Sign.fold_named_context (fun (id,_,_) idl -> id::idl) sign ~init:[]
@@ -721,6 +803,7 @@ let ids_of_context env =
(ids_of_rel_context (rel_context env))
@ (ids_of_named_context (named_context env))
+
let names_of_rel_context env =
List.map (fun (na,_,_) -> na) (rel_context env)
@@ -734,7 +817,6 @@ let rec is_imported_modpath = function
let is_imported_ref = function
| VarRef _ -> false
- | ConstRef kn
| IndRef (kn,_)
| ConstructRef ((kn,_),_)
(* | ModTypeRef ln *) ->
@@ -742,6 +824,8 @@ let is_imported_ref = function
(* | ModRef mp ->
is_imported_modpath mp
*)
+ | ConstRef kn ->
+ let (mp,_,_) = repr_con kn in is_imported_modpath mp
let is_global id =
try
@@ -751,7 +835,7 @@ let is_global id =
false
let is_section_variable id =
- try let _ = Sign.lookup_named id (Global.named_context()) in true
+ try let _ = Global.lookup_named id in true
with Not_found -> false
let next_global_ident_from allow_secvar id avoid =
@@ -861,7 +945,7 @@ let eta_eq_constr =
(* iterator on rel context *)
let process_rel_context f env =
- let sign = named_context env in
+ let sign = named_context_val env in
let rels = rel_context env in
let env0 = reset_with_named_context sign env in
Sign.fold_rel_context f rels ~init:env0
@@ -933,6 +1017,6 @@ let rec rename_bound_var env l c =
| Prod (Anonymous,c1,c2) ->
let env' = push_rel (Anonymous,None,c1) env in
mkProd (Anonymous, c1, rename_bound_var env' l c2)
- | Cast (c,t) -> mkCast (rename_bound_var env l c, t)
+ | Cast (c,k,t) -> mkCast (rename_bound_var env l c, k,t)
| x -> c
diff --git a/pretyping/termops.mli b/pretyping/termops.mli
index 22bd0aba..5f8b5376 100644
--- a/pretyping/termops.mli
+++ b/pretyping/termops.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: termops.mli,v 1.21.2.3 2005/01/21 17:19:37 herbelin Exp $ i*)
+(*i $Id: termops.mli 8003 2006-02-07 22:11:50Z herbelin $ i*)
open Util
open Pp
@@ -16,14 +16,24 @@ open Sign
open Environ
(* Universes *)
-(*i val set_module : Names.dir_path -> unit i*)
val new_univ : unit -> Univ.universe
val new_sort_in_family : sorts_family -> sorts
+val new_Type : unit -> types
+val new_Type_sort : unit -> sorts
+val refresh_universes : types -> types
-(* iterators on terms *)
+(* printers *)
val print_sort : sorts -> std_ppcmds
val print_sort_family : sorts_family -> std_ppcmds
-val print_constr : constr -> std_ppcmds
+(* debug printer: do not use to display terms to the casual user... *)
+val set_print_constr : (env -> constr -> std_ppcmds) -> unit
+val print_constr : constr -> std_ppcmds
+val print_constr_env : env -> constr -> std_ppcmds
+val print_named_context : env -> std_ppcmds
+val print_rel_context : env -> std_ppcmds
+val print_env : env -> std_ppcmds
+
+(* iterators on terms *)
val prod_it : init:types -> (name * types) list -> types
val lam_it : init:constr -> (name * types) list -> constr
val rel_vect : int -> int -> constr array
@@ -43,6 +53,7 @@ val it_named_context_quantifier :
(named_declaration -> 'a -> 'a) -> init:'a -> named_context -> 'a
val it_mkNamedProd_or_LetIn : init:types -> named_context -> types
val it_mkNamedLambda_or_LetIn : init:constr -> named_context -> constr
+val it_mkNamedProd_wo_LetIn : init:types -> named_context -> types
(**********************************************************************)
(* Generic iterators on constr *)
@@ -113,6 +124,8 @@ val subst_term_occ : int list -> constr -> types -> types
val subst_term_occ_decl :
int list -> constr -> named_declaration -> named_declaration
+val error_invalid_occurrence : int list -> 'a
+
(* Alternative term equalities *)
val eta_reduce_head : constr -> constr
val eta_eq_constr : constr -> constr -> bool
@@ -126,8 +139,14 @@ val id_of_name_using_hdchar :
env -> types -> name -> identifier
val named_hd : env -> types -> name -> name
val named_hd_type : env -> types -> name -> name
-val prod_name : env -> name * types * types -> constr
+
+val mkProd_name : env -> name * types * types -> types
+val mkLambda_name : env -> name * types * constr -> constr
+
+(* Deprecated synonyms of [mkProd_name] and [mkLambda_name] *)
+val prod_name : env -> name * types * types -> types
val lambda_name : env -> name * types * constr -> constr
+
val prod_create : env -> types * types -> constr
val lambda_create : env -> types * constr -> constr
val name_assumption : env -> rel_declaration -> rel_declaration
diff --git a/pretyping/typing.ml b/pretyping/typing.ml
index a84cd612..be922c7d 100644
--- a/pretyping/typing.ml
+++ b/pretyping/typing.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: typing.ml,v 1.32.6.2 2004/07/16 19:30:46 herbelin Exp $ *)
+(* $Id: typing.ml 8673 2006-03-29 21:21:52Z herbelin $ *)
open Util
open Names
@@ -16,68 +16,67 @@ open Reductionops
open Type_errors
open Pretype_errors
open Inductive
+open Inductiveops
open Typeops
+open Evd
+
+let meta_type env mv =
+ let ty =
+ try Evd.meta_ftype env mv
+ with Not_found -> error ("unknown meta ?"^string_of_int mv) in
+ meta_instance env ty
let vect_lift = Array.mapi lift
let vect_lift_type = Array.mapi (fun i t -> type_app (lift i) t)
-type 'a mach_flags = {
- fix : bool;
- nocheck : bool }
-
(* The typing machine without information, without universes but with
existential variables. *)
-let assumption_of_judgment env sigma j =
- assumption_of_judgment env (j_nf_evar sigma j)
-
-let type_judgment env sigma j =
- type_judgment env (j_nf_evar sigma j)
-
-
-let rec execute mf env sigma cstr =
+(* cstr must be in n.f. w.r.t. evars and execute returns a judgement
+ where both the term and type are in n.f. *)
+let rec execute env evd cstr =
match kind_of_term cstr with
| Meta n ->
- error "execute: found a non-instanciated goal"
+ { uj_val = cstr; uj_type = nf_evar (evars_of evd) (meta_type evd n) }
| Evar ev ->
- let ty = Instantiate.existential_type sigma ev in
- let jty = execute mf env sigma ty in
- let jty = assumption_of_judgment env sigma jty in
+ let sigma = Evd.evars_of evd in
+ let ty = Evd.existential_type sigma ev in
+ let jty = execute env evd (nf_evar (evars_of evd) ty) in
+ let jty = assumption_of_judgment env jty in
{ uj_val = cstr; uj_type = jty }
| Rel n ->
- judge_of_relative env n
+ j_nf_evar (evars_of evd) (judge_of_relative env n)
| Var id ->
- judge_of_variable env id
+ j_nf_evar (evars_of evd) (judge_of_variable env id)
| Const c ->
- make_judge cstr (constant_type env c)
+ make_judge cstr (nf_evar (evars_of evd) (constant_type env c))
| Ind ind ->
- make_judge cstr (type_of_inductive env ind)
+ make_judge cstr (nf_evar (evars_of evd) (type_of_inductive env ind))
| Construct cstruct ->
- make_judge cstr (type_of_constructor env cstruct)
-
+ make_judge cstr
+ (nf_evar (evars_of evd) (type_of_constructor env cstruct))
+
| Case (ci,p,c,lf) ->
- let cj = execute mf env sigma c in
- let pj = execute mf env sigma p in
- let lfj = execute_array mf env sigma lf in
+ let cj = execute env evd c in
+ let pj = execute env evd p in
+ let lfj = execute_array env evd lf in
let (j,_) = judge_of_case env ci pj cj lfj in
j
| Fix ((vn,i as vni),recdef) ->
- if (not mf.fix) && array_exists (fun n -> n < 0) vn then
- error "General Fixpoints not allowed";
- let (_,tys,_ as recdef') = execute_recdef mf env sigma recdef in
+ let (_,tys,_ as recdef') = execute_recdef env evd recdef in
let fix = (vni,recdef') in
check_fix env fix;
make_judge (mkFix fix) tys.(i)
| CoFix (i,recdef) ->
- let (_,tys,_ as recdef') = execute_recdef mf env sigma recdef in
+ let (_,tys,_ as recdef') = execute_recdef env evd recdef in
let cofix = (i,recdef') in
check_cofix env cofix;
make_judge (mkCoFix cofix) tys.(i)
@@ -89,86 +88,105 @@ let rec execute mf env sigma cstr =
judge_of_type u
| App (f,args) ->
- let j = execute mf env sigma f in
- let jl = execute_array mf env sigma args in
+ let j = execute env evd f in
+ let jl = execute_array env evd args in
let (j,_) = judge_of_apply env j jl in
- j
+ if isInd f then
+ (* Sort-polymorphism of inductive types *)
+ adjust_inductive_level env evd (destInd f) args j
+ else
+ j
| Lambda (name,c1,c2) ->
- let j = execute mf env sigma c1 in
- let var = type_judgment env sigma j in
+ let j = execute env evd c1 in
+ let var = type_judgment env j in
let env1 = push_rel (name,None,var.utj_val) env in
- let j' = execute mf env1 sigma c2 in
+ let j' = execute env1 evd c2 in
judge_of_abstraction env1 name var j'
| Prod (name,c1,c2) ->
- let j = execute mf env sigma c1 in
- let varj = type_judgment env sigma j in
+ let j = execute env evd c1 in
+ let varj = type_judgment env j in
let env1 = push_rel (name,None,varj.utj_val) env in
- let j' = execute mf env1 sigma c2 in
- let varj' = type_judgment env1 sigma j' in
+ let j' = execute env1 evd c2 in
+ let varj' = type_judgment env1 j' in
judge_of_product env name varj varj'
| LetIn (name,c1,c2,c3) ->
- let j1 = execute mf env sigma c1 in
- let j2 = execute mf env sigma c2 in
- let j2 = type_judgment env sigma j2 in
- let _ = judge_of_cast env j1 j2 in
+ let j1 = execute env evd c1 in
+ let j2 = execute env evd c2 in
+ let j2 = type_judgment env j2 in
+ let _ = judge_of_cast env j1 DEFAULTcast j2 in
let env1 = push_rel (name,Some j1.uj_val,j2.utj_val) env in
- let j3 = execute mf env1 sigma c3 in
+ let j3 = execute env1 evd c3 in
judge_of_letin env name j1 j2 j3
- | Cast (c,t) ->
- let cj = execute mf env sigma c in
- let tj = execute mf env sigma t in
- let tj = type_judgment env sigma tj in
- let j, _ = judge_of_cast env cj tj in
+ | Cast (c,k,t) ->
+ let cj = execute env evd c in
+ let tj = execute env evd t in
+ let tj = type_judgment env tj in
+ let j, _ = judge_of_cast env cj k tj in
j
-and execute_recdef mf env sigma (names,lar,vdef) =
- let larj = execute_array mf env sigma lar in
- let lara = Array.map (assumption_of_judgment env sigma) larj in
+and execute_recdef env evd (names,lar,vdef) =
+ let larj = execute_array env evd lar in
+ let lara = Array.map (assumption_of_judgment env) larj in
let env1 = push_rec_types (names,lara,vdef) env in
- let vdefj = execute_array mf env1 sigma vdef in
+ let vdefj = execute_array env1 evd vdef in
let vdefv = Array.map j_val vdefj in
let _ = type_fixpoint env1 names lara vdefj in
(names,lara,vdefv)
-and execute_array mf env sigma v =
- let jl = execute_list mf env sigma (Array.to_list v) in
- Array.of_list jl
-
-and execute_list mf env sigma = function
- | [] ->
- []
- | c::r ->
- let j = execute mf env sigma c in
- let jr = execute_list mf env sigma r in
- j::jr
-
-
-let safe_machine env sigma constr =
- let mf = { fix = false; nocheck = false } in
- execute mf env sigma constr
-
-let unsafe_machine env sigma constr =
- let mf = { fix = false; nocheck = true } in
- execute mf env sigma constr
+and execute_array env evd = Array.map (execute env evd)
+
+and execute_list env evd = List.map (execute env evd)
+
+and adjust_inductive_level env evd ind args j =
+ let specif = lookup_mind_specif env ind in
+ if is_small_inductive specif then
+ (* No polymorphism *)
+ j
+ else
+ (* Retyping constructor with the actual arguments *)
+ let env',llc,ls0 = constructor_instances env specif ind args in
+ let llj = Array.map (execute_array env' evd) llc in
+ let ls =
+ Array.map (fun lj ->
+ let ls =
+ Array.map (fun c -> decomp_sort env (evars_of evd) c.uj_type) lj
+ in
+ max_inductive_sort ls) llj
+ in
+ let s = find_inductive_level env specif ind ls0 ls in
+ on_judgment_type (set_inductive_level env s) j
+
+let mcheck env evd c t =
+ let sigma = Evd.evars_of evd in
+ let j = execute env evd (nf_evar sigma c) in
+ if not (is_conv_leq env sigma j.uj_type t) then
+ error_actual_type env j (nf_evar sigma t)
(* Type of a constr *)
-
-let type_of env sigma c =
- let j = safe_machine env sigma c in
+
+let mtype_of env evd c =
+ let j = execute env evd (nf_evar (evars_of evd) c) in
(* No normalization: it breaks Pattern! *)
(*nf_betaiota*) (body_of_type j.uj_type)
-(* The typed type of a judgment. *)
+let msort_of env evd c =
+ let j = execute env evd (nf_evar (evars_of evd) c) in
+ let a = type_judgment env j in
+ a.utj_type
-let execute_type env sigma constr =
- let j = execute { fix=false; nocheck=true } env sigma constr in
- assumption_of_judgment env sigma j
+let type_of env sigma c =
+ mtype_of env (Evd.create_evar_defs sigma) c
+let sort_of env sigma c =
+ msort_of env (Evd.create_evar_defs sigma) c
+let check env sigma c =
+ mcheck env (Evd.create_evar_defs sigma) c
-let execute_rec_type env sigma constr =
- let j = execute { fix=false; nocheck=false } env sigma constr in
- assumption_of_judgment env sigma j
+(* The typed type of a judgment. *)
+let mtype_of_type env evd constr =
+ let j = execute env evd (nf_evar (evars_of evd) constr) in
+ assumption_of_judgment env j
diff --git a/pretyping/typing.mli b/pretyping/typing.mli
index 4ea74dcd..c9d7d572 100644
--- a/pretyping/typing.mli
+++ b/pretyping/typing.mli
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i $Id: typing.mli,v 1.7.14.1 2004/07/16 19:30:47 herbelin Exp $ i*)
+(*i $Id: typing.mli 6113 2004-09-17 20:28:19Z barras $ i*)
(*i*)
open Term
@@ -17,11 +17,18 @@ open Evd
(* This module provides the typing machine with existential variables
(but without universes). *)
-val unsafe_machine : env -> evar_map -> constr -> unsafe_judgment
-
-val type_of : env -> evar_map -> constr -> constr
-
-val execute_type : env -> evar_map -> constr -> types
-
-val execute_rec_type : env -> evar_map -> constr -> types
-
+(* Typecheck a term and return its type *)
+val type_of : env -> evar_map -> constr -> types
+(* Typecheck a type and return its sort *)
+val sort_of : env -> evar_map -> types -> sorts
+(* Typecheck a term has a given type (assuming the type is OK *)
+val check : env -> evar_map -> constr -> types -> unit
+
+(* The same but with metas... *)
+val mtype_of : env -> evar_defs -> constr -> types
+val msort_of : env -> evar_defs -> types -> sorts
+val mcheck : env -> evar_defs -> constr -> types -> unit
+val meta_type : evar_defs -> metavariable -> types
+
+(* unused typing function... *)
+val mtype_of_type : env -> evar_defs -> types -> types
diff --git a/pretyping/unification.ml b/pretyping/unification.ml
new file mode 100644
index 00000000..e51f5e0e
--- /dev/null
+++ b/pretyping/unification.ml
@@ -0,0 +1,471 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(* $Id: unification.ml 7113 2005-06-05 17:13:06Z barras $ *)
+
+open Pp
+open Util
+open Names
+open Nameops
+open Term
+open Termops
+open Sign
+open Environ
+open Evd
+open Reduction
+open Reductionops
+open Rawterm
+open Pattern
+open Evarutil
+open Pretype_errors
+
+(* if lname_typ is [xn,An;..;x1,A1] and l is a list of terms,
+ gives [x1:A1]..[xn:An]c' such that c converts to ([x1:A1]..[xn:An]c' l) *)
+
+let abstract_scheme env c l lname_typ =
+ List.fold_left2
+ (fun t (locc,a) (na,_,ta) ->
+ let na = match kind_of_term a with Var id -> Name id | _ -> na in
+ if occur_meta ta then error "cannot find a type for the generalisation"
+ else if occur_meta a then lambda_name env (na,ta,t)
+ else lambda_name env (na,ta,subst_term_occ locc a t))
+ c
+ (List.rev l)
+ lname_typ
+
+let abstract_list_all env sigma typ c l =
+ let ctxt,_ = decomp_n_prod env sigma (List.length l) typ in
+ let p = abstract_scheme env c (List.map (function a -> [],a) l) ctxt in
+ try
+ if is_conv_leq env sigma (Typing.type_of env sigma p) typ then p
+ else error "abstract_list_all"
+ with UserError _ ->
+ raise (PretypeError (env,CannotGeneralize typ))
+
+
+(*******************************)
+
+(* Unification à l'ordre 0 de m et n: [unify_0 env sigma cv_pb m n]
+ renvoie deux listes:
+
+ metasubst:(int*constr)list récolte les instances des (Meta k)
+ evarsubst:(constr*constr)list récolte les instances des (Const "?k")
+
+ Attention : pas d'unification entre les différences instances d'une
+ même meta ou evar, il peut rester des doublons *)
+
+(* Unification order: *)
+(* Left to right: unifies first argument and then the other arguments *)
+(*let unify_l2r x = List.rev x
+(* Right to left: unifies last argument and then the other arguments *)
+let unify_r2l x = x
+
+let sort_eqns = unify_r2l
+*)
+
+let unify_0 env sigma cv_pb mod_delta m n =
+ let trivial_unify pb substn m n =
+ if (not(occur_meta m)) && (if mod_delta then is_fconv pb env sigma m n else eq_constr m n) then substn
+ else error_cannot_unify env sigma (m,n) in
+ let rec unirec_rec pb ((metasubst,evarsubst) as substn) m n =
+ let cM = Evarutil.whd_castappevar sigma m
+ and cN = Evarutil.whd_castappevar sigma n in
+ match (kind_of_term cM,kind_of_term cN) with
+ | Meta k1, Meta k2 ->
+ if k1 < k2 then (k1,cN)::metasubst,evarsubst
+ else if k1 = k2 then substn
+ else (k2,cM)::metasubst,evarsubst
+ | Meta k, _ -> (k,cN)::metasubst,evarsubst
+ | _, Meta k -> (k,cM)::metasubst,evarsubst
+ | Evar _, _ -> metasubst,((cM,cN)::evarsubst)
+ | _, Evar _ -> metasubst,((cN,cM)::evarsubst)
+
+ | Lambda (_,t1,c1), Lambda (_,t2,c2) ->
+ unirec_rec CONV (unirec_rec CONV substn t1 t2) c1 c2
+ | Prod (_,t1,c1), Prod (_,t2,c2) ->
+ unirec_rec pb (unirec_rec CONV substn t1 t2) c1 c2
+ | LetIn (_,b,_,c), _ -> unirec_rec pb substn (subst1 b c) cN
+ | _, LetIn (_,b,_,c) -> unirec_rec pb substn cM (subst1 b c)
+
+ | App (f1,l1), App (f2,l2) ->
+ let len1 = Array.length l1
+ and len2 = Array.length l2 in
+ let (f1,l1,f2,l2) =
+ if len1 = len2 then (f1,l1,f2,l2)
+ else if len1 < len2 then
+ let extras,restl2 = array_chop (len2-len1) l2 in
+ (f1, l1, appvect (f2,extras), restl2)
+ else
+ let extras,restl1 = array_chop (len1-len2) l1 in
+ (appvect (f1,extras), restl1, f2, l2) in
+ (try
+ array_fold_left2 (unirec_rec CONV)
+ (unirec_rec CONV substn f1 f2) l1 l2
+ with ex when precatchable_exception ex ->
+ trivial_unify pb substn cM cN)
+ | Case (_,p1,c1,cl1), Case (_,p2,c2,cl2) ->
+ array_fold_left2 (unirec_rec CONV)
+ (unirec_rec CONV (unirec_rec CONV substn p1 p2) c1 c2) cl1 cl2
+
+ | _ -> trivial_unify pb substn cM cN
+
+ in
+ if (not(occur_meta m)) &&
+ (if mod_delta then is_fconv cv_pb env sigma m n else eq_constr m n)
+ then
+ ([],[])
+ else
+ let (mc,ec) = unirec_rec cv_pb ([],[]) m n in
+ ((*sort_eqns*) mc, (*sort_eqns*) ec)
+
+
+(* Unification
+ *
+ * Procedure:
+ * (1) The function [unify mc wc M N] produces two lists:
+ * (a) a list of bindings Meta->RHS
+ * (b) a list of bindings EVAR->RHS
+ *
+ * The Meta->RHS bindings cannot themselves contain
+ * meta-vars, so they get applied eagerly to the other
+ * bindings. This may or may not close off all RHSs of
+ * the EVARs. For each EVAR whose RHS is closed off,
+ * we can just apply it, and go on. For each which
+ * is not closed off, we need to do a mimick step -
+ * in general, we have something like:
+ *
+ * ?X == (c e1 e2 ... ei[Meta(k)] ... en)
+ *
+ * so we need to do a mimick step, converting ?X
+ * into
+ *
+ * ?X -> (c ?z1 ... ?zn)
+ *
+ * of the proper types. Then, we can decompose the
+ * equation into
+ *
+ * ?z1 --> e1
+ * ...
+ * ?zi --> ei[Meta(k)]
+ * ...
+ * ?zn --> en
+ *
+ * and keep on going. Whenever we find that a R.H.S.
+ * is closed, we can, as before, apply the constraint
+ * directly. Whenever we find an equation of the form:
+ *
+ * ?z -> Meta(n)
+ *
+ * we can reverse the equation, put it into our metavar
+ * substitution, and keep going.
+ *
+ * The most efficient mimick possible is, for each
+ * Meta-var remaining in the term, to declare a
+ * new EVAR of the same type. This is supposedly
+ * determinable from the clausale form context -
+ * we look up the metavar, take its type there,
+ * and apply the metavar substitution to it, to
+ * close it off. But this might not always work,
+ * since other metavars might also need to be resolved. *)
+
+let applyHead env evd n c =
+ let rec apprec n c cty evd =
+ if n = 0 then
+ (evd, c)
+ else
+ match kind_of_term (whd_betadeltaiota env (evars_of evd) cty) with
+ | Prod (_,c1,c2) ->
+ let (evd',evar) = Evarutil.new_evar evd env c1 in
+ apprec (n-1) (mkApp(c,[|evar|])) (subst1 evar c2) evd'
+ | _ -> error "Apply_Head_Then"
+ in
+ apprec n c (Typing.type_of env (evars_of evd) c) evd
+
+let is_mimick_head f =
+ match kind_of_term f with
+ (Const _|Var _|Rel _|Construct _|Ind _) -> true
+ | _ -> false
+
+(* [w_merge env sigma b metas evars] merges common instances in metas
+ or in evars, possibly generating new unification problems; if [b]
+ is true, unification of types of metas is required *)
+
+let w_merge env with_types mod_delta metas evars evd =
+ let ty_metas = ref [] in
+ let ty_evars = ref [] in
+ let rec w_merge_rec evd metas evars =
+ match (evars,metas) with
+ | ([], []) -> evd
+
+ | ((lhs,rhs)::t, metas) ->
+ (match kind_of_term rhs with
+
+ | Meta k -> w_merge_rec evd ((k,lhs)::metas) t
+
+ | krhs ->
+ (match kind_of_term lhs with
+
+ | Evar (evn,_ as ev) ->
+ if is_defined_evar evd ev then
+ let (metas',evars') =
+ unify_0 env (evars_of evd) CONV mod_delta rhs lhs in
+ w_merge_rec evd (metas'@metas) (evars'@t)
+ else begin
+ let rhs' =
+ if occur_meta rhs then subst_meta metas rhs else rhs
+ in
+ if occur_evar evn rhs' then
+ error "w_merge: recursive equation";
+ match krhs with
+ | App (f,cl) when is_mimick_head f ->
+ (try
+ w_merge_rec (fst (evar_define env ev rhs' evd)) metas t
+ with ex when precatchable_exception ex ->
+ let evd' =
+ mimick_evar evd mod_delta f (Array.length cl) evn in
+ w_merge_rec evd' metas evars)
+ | _ ->
+ (* ensure tail recursion in non-mimickable case! *)
+ w_merge_rec (fst (evar_define env ev rhs' evd)) metas t
+ end
+
+ | _ -> anomaly "w_merge_rec"))
+
+ | ([], (mv,n)::t) ->
+ if meta_defined evd mv then
+ let (metas',evars') =
+ unify_0 env (evars_of evd) CONV mod_delta
+ (meta_fvalue evd mv).rebus n in
+ w_merge_rec evd (metas'@t) evars'
+ else
+ begin
+ if with_types (* or occur_meta mvty *) then
+ (let mvty = Typing.meta_type evd mv in
+ try
+ let sigma = evars_of evd in
+ (* why not typing with the metamap ? *)
+ let nty = Typing.type_of env sigma (nf_meta evd n) in
+ let (mc,ec) = unify_0 env sigma CUMUL mod_delta nty mvty in
+ ty_metas := mc @ !ty_metas;
+ ty_evars := ec @ !ty_evars
+ with e when precatchable_exception e -> ());
+ let evd' = meta_assign mv n evd in
+ w_merge_rec evd' t []
+ end
+
+ and mimick_evar evd mod_delta hdc nargs sp =
+ let ev = Evd.map (evars_of evd) sp in
+ let sp_env = Global.env_of_context ev.evar_hyps in
+ let (evd', c) = applyHead sp_env evd nargs hdc in
+ let (mc,ec) =
+ unify_0 sp_env (evars_of evd') CUMUL mod_delta
+ (Retyping.get_type_of sp_env (evars_of evd') c) ev.evar_concl in
+ let evd'' = w_merge_rec evd' mc ec in
+ if (evars_of evd') == (evars_of evd'')
+ then Evd.evar_define sp c evd''
+ else Evd.evar_define sp (Evarutil.nf_evar (evars_of evd'') c) evd'' in
+
+ (* merge constraints *)
+ let evd' = w_merge_rec evd metas evars in
+ if with_types then
+ (* merge constraints about types: if they fail, don't worry *)
+ try w_merge_rec evd' !ty_metas !ty_evars
+ with e when precatchable_exception e -> evd'
+ else
+ evd'
+
+(* [w_unify env evd M N]
+ performs a unification of M and N, generating a bunch of
+ unification constraints in the process. These constraints
+ are processed, one-by-one - they may either generate new
+ bindings, or, if there is already a binding, new unifications,
+ which themselves generate new constraints. This continues
+ until we get failure, or we run out of constraints.
+ [clenv_typed_unify M N clenv] expects in addition that expected
+ types of metavars are unifiable with the types of their instances *)
+
+let w_unify_core_0 env with_types cv_pb mod_delta m n evd =
+ let (mc,ec) = unify_0 env (evars_of evd) cv_pb mod_delta m n in
+ w_merge env with_types mod_delta mc ec evd
+
+let w_unify_0 env = w_unify_core_0 env false
+let w_typed_unify env = w_unify_core_0 env true
+
+
+(* takes a substitution s, an open term op and a closed term cl
+ try to find a subterm of cl which matches op, if op is just a Meta
+ FAIL because we cannot find a binding *)
+
+let iter_fail f a =
+ let n = Array.length a in
+ let rec ffail i =
+ if i = n then error "iter_fail"
+ else
+ try f a.(i)
+ with ex when precatchable_exception ex -> ffail (i+1)
+ in ffail 0
+
+(* Tries to find an instance of term [cl] in term [op].
+ Unifies [cl] to every subterm of [op] until it finds a match.
+ Fails if no match is found *)
+let w_unify_to_subterm env ?(mod_delta=true) (op,cl) evd =
+ let rec matchrec cl =
+ let cl = strip_outer_cast cl in
+ (try
+ if closed0 cl
+ then w_unify_0 env CONV mod_delta op cl evd,cl
+ else error "Bound 1"
+ with ex when precatchable_exception ex ->
+ (match kind_of_term cl with
+ | App (f,args) ->
+ let n = Array.length args in
+ assert (n>0);
+ let c1 = mkApp (f,Array.sub args 0 (n-1)) in
+ let c2 = args.(n-1) in
+ (try
+ matchrec c1
+ with ex when precatchable_exception ex ->
+ matchrec c2)
+ | Case(_,_,c,lf) -> (* does not search in the predicate *)
+ (try
+ matchrec c
+ with ex when precatchable_exception ex ->
+ iter_fail matchrec lf)
+ | LetIn(_,c1,_,c2) ->
+ (try
+ matchrec c1
+ with ex when precatchable_exception ex ->
+ matchrec c2)
+
+ | Fix(_,(_,types,terms)) ->
+ (try
+ iter_fail matchrec types
+ with ex when precatchable_exception ex ->
+ iter_fail matchrec terms)
+
+ | CoFix(_,(_,types,terms)) ->
+ (try
+ iter_fail matchrec types
+ with ex when precatchable_exception ex ->
+ iter_fail matchrec terms)
+
+ | Prod (_,t,c) ->
+ (try
+ matchrec t
+ with ex when precatchable_exception ex ->
+ matchrec c)
+ | Lambda (_,t,c) ->
+ (try
+ matchrec t
+ with ex when precatchable_exception ex ->
+ matchrec c)
+ | _ -> error "Match_subterm"))
+ in
+ try matchrec cl
+ with ex when precatchable_exception ex ->
+ raise (PretypeError (env,NoOccurrenceFound op))
+
+let w_unify_to_subterm_list env mod_delta allow_K oplist t evd =
+ List.fold_right
+ (fun op (evd,l) ->
+ if isMeta op then
+ if allow_K then (evd,op::l)
+ else error "Match_subterm"
+ else if occur_meta op then
+ let (evd',cl) =
+ try
+ (* This is up to delta for subterms w/o metas ... *)
+ w_unify_to_subterm env ~mod_delta (strip_outer_cast op,t) evd
+ with PretypeError (env,NoOccurrenceFound _) when allow_K -> (evd,op)
+ in
+ (evd',cl::l)
+ else if allow_K or dependent op t then
+ (evd,op::l)
+ else
+ (* This is not up to delta ... *)
+ raise (PretypeError (env,NoOccurrenceFound op)))
+ oplist
+ (evd,[])
+
+let secondOrderAbstraction env mod_delta allow_K typ (p, oplist) evd =
+ let sigma = evars_of evd in
+ let (evd',cllist) =
+ w_unify_to_subterm_list env mod_delta allow_K oplist typ evd in
+ let typp = Typing.meta_type evd' p in
+ let pred = abstract_list_all env sigma typp typ cllist in
+ w_unify_0 env CONV mod_delta (mkMeta p) pred evd'
+
+let w_unify2 env mod_delta allow_K cv_pb ty1 ty2 evd =
+ let c1, oplist1 = whd_stack ty1 in
+ let c2, oplist2 = whd_stack ty2 in
+ match kind_of_term c1, kind_of_term c2 with
+ | Meta p1, _ ->
+ (* Find the predicate *)
+ let evd' =
+ secondOrderAbstraction env mod_delta allow_K ty2 (p1,oplist1) evd in
+ (* Resume first order unification *)
+ w_unify_0 env cv_pb mod_delta (nf_meta evd' ty1) ty2 evd'
+ | _, Meta p2 ->
+ (* Find the predicate *)
+ let evd' =
+ secondOrderAbstraction env mod_delta allow_K ty1 (p2, oplist2) evd in
+ (* Resume first order unification *)
+ w_unify_0 env cv_pb mod_delta ty1 (nf_meta evd' ty2) evd'
+ | _ -> error "w_unify2"
+
+
+(* The unique unification algorithm works like this: If the pattern is
+ flexible, and the goal has a lambda-abstraction at the head, then
+ we do a first-order unification.
+
+ If the pattern is not flexible, then we do a first-order
+ unification, too.
+
+ If the pattern is flexible, and the goal doesn't have a
+ lambda-abstraction head, then we second-order unification. *)
+
+(* We decide here if first-order or second-order unif is used for Apply *)
+(* We apply a term of type (ai:Ai)C and try to solve a goal C' *)
+(* The type C is in clenv.templtyp.rebus with a lot of Meta to solve *)
+
+(* 3-4-99 [HH] New fo/so choice heuristic :
+ In case we have to unify (Meta(1) args) with ([x:A]t args')
+ we first try second-order unification and if it fails first-order.
+ Before, second-order was used if the type of Meta(1) and [x:A]t was
+ convertible and first-order otherwise. But if failed if e.g. the type of
+ Meta(1) had meta-variables in it. *)
+let w_unify allow_K env cv_pb ?(mod_delta=true) ty1 ty2 evd =
+ let hd1,l1 = whd_stack ty1 in
+ let hd2,l2 = whd_stack ty2 in
+ match kind_of_term hd1, l1<>[], kind_of_term hd2, l2<>[] with
+ (* Pattern case *)
+ | (Meta _, true, Lambda _, _ | Lambda _, _, Meta _, true)
+ when List.length l1 = List.length l2 ->
+ (try
+ w_typed_unify env cv_pb mod_delta ty1 ty2 evd
+ with ex when precatchable_exception ex ->
+ try
+ w_unify2 env mod_delta allow_K cv_pb ty1 ty2 evd
+ with PretypeError (env,NoOccurrenceFound c) as e -> raise e
+ | ex when precatchable_exception ex ->
+ error "Cannot solve a second-order unification problem")
+
+ (* Second order case *)
+ | (Meta _, true, _, _ | _, _, Meta _, true) ->
+ (try
+ w_unify2 env mod_delta allow_K cv_pb ty1 ty2 evd
+ with PretypeError (env,NoOccurrenceFound c) as e -> raise e
+ | ex when precatchable_exception ex ->
+ try
+ w_typed_unify env cv_pb mod_delta ty1 ty2 evd
+ with ex when precatchable_exception ex ->
+ error "Cannot solve a second-order unification problem")
+
+ (* General case: try first order *)
+ | _ -> w_unify_0 env cv_pb mod_delta ty1 ty2 evd
+
diff --git a/pretyping/unification.mli b/pretyping/unification.mli
new file mode 100644
index 00000000..6be530be
--- /dev/null
+++ b/pretyping/unification.mli
@@ -0,0 +1,33 @@
+(************************************************************************)
+(* v * The Coq Proof Assistant / The Coq Development Team *)
+(* <O___,, * CNRS-Ecole Polytechnique-INRIA Futurs-Universite Paris Sud *)
+(* \VV/ **************************************************************)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(************************************************************************)
+
+(*i $Id: unification.mli 6142 2004-09-27 19:33:01Z sacerdot $ i*)
+
+(*i*)
+open Term
+open Environ
+open Evd
+(*i*)
+
+(* The "unique" unification fonction *)
+val w_unify :
+ bool -> env -> conv_pb -> ?mod_delta:bool -> constr -> constr -> evar_defs -> evar_defs
+
+(* [w_unify_to_subterm env (c,t) m] performs unification of [c] with a
+ subterm of [t]. Constraints are added to [m] and the matched
+ subterm of [t] is also returned. *)
+val w_unify_to_subterm :
+ env -> ?mod_delta:bool -> constr * constr -> evar_defs -> evar_defs * constr
+
+(*i This should be in another module i*)
+
+(* [abstract_list_all env sigma t c l] *)
+(* abstracts the terms in l over c to get a term of type t *)
+(* (exported for inv.ml) *)
+val abstract_list_all :
+ env -> evar_map -> constr -> constr -> constr list -> constr