summaryrefslogtreecommitdiff
path: root/pretyping/pretyping.ml
diff options
context:
space:
mode:
Diffstat (limited to 'pretyping/pretyping.ml')
-rw-r--r--pretyping/pretyping.ml421
1 files changed, 229 insertions, 192 deletions
diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml
index 0db64a52..5f0999cb 100644
--- a/pretyping/pretyping.ml
+++ b/pretyping/pretyping.ml
@@ -6,7 +6,7 @@
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(* $Id: pretyping.ml 9976 2007-07-12 11:58:30Z msozeau $ *)
+(* $Id: pretyping.ml 11047 2008-06-03 23:08:00Z msozeau $ *)
open Pp
open Util
@@ -43,6 +43,34 @@ open Inductiveops
(************************************************************************)
+(* An auxiliary function for searching for fixpoint guard indexes *)
+
+exception Found of int array
+
+let search_guard loc env possible_indexes fixdefs =
+ (* Standard situation with only one possibility for each fix. *)
+ (* We treat it separately in order to get proper error msg. *)
+ if List.for_all (fun l->1=List.length l) possible_indexes then
+ let indexes = Array.of_list (List.map List.hd possible_indexes) in
+ let fix = ((indexes, 0),fixdefs) in
+ (try check_fix env fix with
+ | e -> if loc = dummy_loc then raise e else Stdpp.raise_with_loc loc e);
+ indexes
+ else
+ (* we now search recursively amoungst all combinations *)
+ (try
+ List.iter
+ (fun l ->
+ let indexes = Array.of_list l in
+ let fix = ((indexes, 0),fixdefs) in
+ try check_fix env fix; raise (Found indexes)
+ with TypeError _ -> ())
+ (list_combinations possible_indexes);
+ let errmsg = "cannot guess decreasing argument of fix" in
+ if loc = dummy_loc then error errmsg else
+ user_err_loc (loc,"search_guard", Pp.str errmsg)
+ with Found indexes -> indexes)
+
(* To embed constr in rawconstr *)
let ((constr_in : constr -> Dyn.t),
(constr_out : Dyn.t -> constr)) = create "constr"
@@ -70,7 +98,7 @@ sig
unresolved holes as evars and returning the typing contexts of
these evars. Work as [understand_gen] for the rest. *)
- val understand_tcc :
+ val understand_tcc : ?resolve_classes:bool ->
evar_map -> env -> ?expected_type:types -> rawconstr -> open_constr
val understand_tcc_evars :
@@ -129,7 +157,7 @@ sig
rawconstr -> unsafe_type_judgment
val pretype_gen :
- evar_defs ref -> env ->
+ evar_defs ref -> env ->
var_map * (identifier * identifier option) list ->
typing_constraint -> rawconstr -> constr
@@ -143,90 +171,83 @@ module Pretyping_F (Coercion : Coercion.S) = struct
(* 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';
+ let evd_comb0 f evdref =
+ let (evd',x) = f !evdref in
+ evdref := evd';
x
- let evd_comb1 f isevars x =
- let (evd',y) = f !isevars x in
- isevars := evd';
+ let evd_comb1 f evdref x =
+ let (evd',y) = f !evdref x in
+ evdref := evd';
y
- let evd_comb2 f isevars x y =
- let (evd',z) = f !isevars x y in
- isevars := evd';
+ let evd_comb2 f evdref x y =
+ let (evd',z) = f !evdref x y in
+ evdref := evd';
z
- let evd_comb3 f isevars x y z =
- let (evd',t) = f !isevars x y z in
- isevars := evd';
+ let evd_comb3 f evdref x y z =
+ let (evd',t) = f !evdref x y z in
+ evdref := 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 evar_type_fixpoint loc env evdref 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
+ if not (e_cumul env evdref (vdefj.(i)).uj_type
(lift lt lar.(i))) then
- error_ill_typed_rec_body_loc loc env (evars_of !isevars)
+ error_ill_typed_rec_body_loc loc env (evars_of !evdref)
i lna vdefj lar
done
- let check_branches_message loc env isevars c (explft,lft) =
+ let check_branches_message loc env evdref 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
+ if not (e_cumul env evdref lft.(i) explft.(i)) then
+ let sigma = evars_of !evdref 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
+ let inh_conv_coerce_to_tycon loc env evdref j = function
| None -> j
- | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to loc env) isevars j t
+ | Some t -> evd_comb2 (Coercion.inh_conv_coerce_to loc env) evdref 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)
- *)
+ (* used to enforce a name in Lambda when the type constraints itself
+ is named, hence possibly dependent *)
- 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 orelse_name name name' = match name with
+ | Anonymous -> name'
+ | _ -> name
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 "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
+ try
+ let (n,typ) = lookup_rel_id id (rel_context env) in
+ { uj_val = mkRel n; uj_type = 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 "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 *)
@@ -251,7 +272,7 @@ module Pretyping_F (Coercion : Coercion.S) = struct
(*************************************************************************)
(* Main pretyping function *)
- let pretype_ref isevars env ref =
+ let pretype_ref evdref env ref =
let c = constr_of_global ref in
make_judge c (Retyping.get_type_of env Evd.empty c)
@@ -259,30 +280,32 @@ module Pretyping_F (Coercion : Coercion.S) = struct
| 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 *)
+ exception Found of fixpoint
+
+ (* [pretype tycon env evdref lvar lmeta cstr] attempts to type [cstr] *)
+ (* in environment [env], with existential variables [evdref] and *)
(* the type constraint tycon *)
- let rec pretype (tycon : type_constraint) env isevars lvar = function
+ let rec pretype (tycon : type_constraint) env evdref lvar = function
| RRef (loc,ref) ->
- inh_conv_coerce_to_tycon loc env isevars
- (pretype_ref isevars env ref)
+ inh_conv_coerce_to_tycon loc env evdref
+ (pretype_ref evdref env ref)
tycon
| RVar (loc, id) ->
- inh_conv_coerce_to_tycon loc env isevars
+ inh_conv_coerce_to_tycon loc env evdref
(pretype_id loc env lvar id)
tycon
- | REvar (loc, ev, instopt) ->
+ | REvar (loc, evk, 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.find (evars_of !isevars) ev) in
+ let hyps = evar_filtered_context (Evd.find (evars_of !evdref) evk) 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
+ let c = mkEvar (evk, args) in
+ let j = (Retyping.get_judgment_of env (evars_of !evdref) c) in
+ inh_conv_coerce_to_tycon loc env evdref j tycon
| RPatVar (loc,(someta,n)) ->
anomaly "Found a pattern variable in a rawterm to type"
@@ -292,26 +315,26 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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 }
+ e_new_evar evdref env ~src:(loc,InternalHole) (new_Type ()) in
+ { uj_val = e_new_evar evdref 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
+ | (na,bk,None,ty)::bl ->
+ let ty' = pretype_type empty_valcon env evdref 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
+ | (na,bk,Some bd,ty)::bl ->
+ let ty' = pretype_type empty_valcon env evdref lvar ty in
+ let bd' = pretype (mk_tycon ty'.utj_val) env evdref 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)
+ pretype_type empty_valcon (push_rel_context e env) evdref 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
@@ -328,115 +351,111 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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
+ let j = pretype (mk_tycon ty) nenv evdref 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;
+ evar_type_fixpoint loc env evdref names ftys vdefj;
let fixj = match fixkind with
| RFix (vn,i) ->
- let guard_indexes = Array.mapi
+ (* First, let's find the guard indexes. *)
+ (* If recursive argument was not given by user, we try all args.
+ An earlier approach was to look only for inductive arguments,
+ but doing it properly involves delta-reduction, and it finally
+ doesn't seem worth the effort (except for huge mutual
+ fixpoints ?) *)
+ let possible_indexes = Array.to_list (Array.mapi
(fun i (n,_) -> match n with
- | Some n -> n
- | None ->
- (* Recursive argument was not given by the user : We
- check that there is only one inductive argument *)
- let ctx = ctxtv.(i) in
- let isIndApp t =
- isInd (fst (decompose_app (strip_head_cast t))) in
- (* This could be more precise (e.g. do some delta) *)
- let lb = List.rev_map (fun (_,_,t) -> isIndApp t) ctx in
- try (list_unique_index true lb) - 1
- with Not_found ->
- Util.user_err_loc
- (loc,"pretype",
- Pp.str "cannot guess decreasing argument of fix"))
- vn
- in
- let fix = ((guard_indexes, 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 ->
+ | Some n -> [n]
+ | None -> list_map_i (fun i _ -> i) 0 ctxtv.(i))
+ vn)
+ in
+ let fixdecls = (names,ftys,Array.map j_val vdefj) in
+ let indexes = search_guard loc env possible_indexes fixdecls in
+ make_judge (mkFix ((indexes,i),fixdecls)) 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
+ inh_conv_coerce_to_tycon loc env evdref fixj tycon
| RSort (loc,s) ->
- inh_conv_coerce_to_tycon loc env isevars (pretype_sort s) tycon
+ inh_conv_coerce_to_tycon loc env evdref (pretype_sort s) tycon
| RApp (loc,f,args) ->
- let fj = pretype empty_tycon env isevars lvar f in
+ let fj = pretype empty_tycon env evdref 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 = evd_comb1 (Coercion.inh_app_fun env) isevars resj in
- let resty = whd_betadeltaiota env (evars_of !isevars) resj.uj_type in
+ let resj = evd_comb1 (Coercion.inh_app_fun env) evdref resj in
+ let resty = whd_betadeltaiota env (evars_of !evdref) 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 hj = pretype (mk_tycon c1) env evdref lvar c in
let value, typ = applist (j_val resj, [j_val hj]), subst1 hj.uj_val c2 in
apply_rec env (n+1)
{ uj_val = value;
uj_type = typ }
rest
| _ ->
- let hj = pretype empty_tycon env isevars lvar c in
+ let hj = pretype empty_tycon env evdref lvar c in
error_cant_apply_not_functional_loc
- (join_loc floc argloc) env (evars_of !isevars)
+ (join_loc floc argloc) env (evars_of !evdref)
resj [hj]
in
let resj = apply_rec env 1 fj args in
let resj =
- match evar_kind_of_term !isevars resj.uj_val with
+ match evar_kind_of_term !evdref resj.uj_val with
| App (f,args) ->
- let f = whd_evar (Evd.evars_of !isevars) f in
+ let f = whd_evar (Evd.evars_of !evdref) f in
begin match kind_of_term f with
- | Ind _ (* | Const _ *) ->
- let sigma = evars_of !isevars in
+ | Ind _ | Const _
+ when isInd f or has_polymorphic_type (destConst f)
+ ->
+ let sigma = evars_of !evdref in
let c = mkApp (f,Array.map (whd_evar sigma) args) in
let t = Retyping.get_type_of env sigma c in
make_judge c t
| _ -> resj end
| _ -> resj in
- inh_conv_coerce_to_tycon loc env isevars resj tycon
+ inh_conv_coerce_to_tycon loc env evdref resj tycon
- | RLambda(loc,name,c1,c2) ->
- let (name',dom,rng) = evd_comb1 (split_tycon loc env) isevars tycon in
+ | RLambda(loc,name,bk,c1,c2) ->
+ let (name',dom,rng) = evd_comb1 (split_tycon loc env) evdref tycon in
let dom_valcon = valcon_of_tycon dom in
- let j = pretype_type dom_valcon env isevars lvar c1 in
+ let j = pretype_type dom_valcon env evdref 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'
+ let j' = pretype rng (push_rel var env) evdref lvar c2 in
+ judge_of_abstraction env (orelse_name name name') j j'
- | RProd(loc,name,c1,c2) ->
- let j = pretype_type empty_valcon env isevars lvar c1 in
+ | RProd(loc,name,bk,c1,c2) ->
+ let j = pretype_type empty_valcon env evdref 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 j' = pretype_type empty_valcon env' evdref 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
+ inh_conv_coerce_to_tycon loc env evdref resj tycon
| RLetIn(loc,name,c1,c2) ->
- let j = pretype empty_tycon env isevars lvar c1 in
+ let j = pretype empty_tycon env evdref 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
+ let j' = pretype tycon (push_rel var env) evdref 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 cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
- try find_rectype env (evars_of !isevars) cj.uj_type
+ try find_rectype env (evars_of !evdref) cj.uj_type
with Not_found ->
let cloc = loc_of_rawconstr c in
- error_case_not_inductive_loc cloc env (evars_of !isevars) cj
+ error_case_not_inductive_loc cloc env (evars_of !evdref) cj
in
let cstrs = get_constructors env indf in
if Array.length cstrs <> 1 then
@@ -459,50 +478,50 @@ module Pretyping_F (Coercion : Coercion.S) = struct
(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_isevar !isevars pj.utj_val in
+ let pj = pretype_type empty_valcon env_p evdref lvar p in
+ let ccl = nf_isevar !evdref 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 fty = hnf_lam_applist env (evars_of !evdref) lp inst in
+ let fj = pretype (mk_tycon fty) env_f evdref 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
+ let ci = make_case_info env mis LetStyle 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 fj = pretype tycon env_f evdref lvar d in
let f = it_mkLambda_or_LetIn fj.uj_val fsign in
- let ccl = nf_isevar !isevars fj.uj_type in
+ let ccl = nf_isevar !evdref 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)
+ error_cant_find_case_type_loc loc env (evars_of !evdref)
cj.uj_val in
let ccl = refresh_universes ccl 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
+ let ci = make_case_info env mis LetStyle 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 cj = pretype empty_tycon env evdref lvar c in
let (IndType (indf,realargs)) =
- try find_rectype env (evars_of !isevars) cj.uj_type
+ try find_rectype env (evars_of !evdref) 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
+ error_case_not_inductive_loc cloc env (evars_of !evdref) cj in
let cstrs = get_constructors env indf in
if Array.length cstrs <> 2 then
user_err_loc (loc,"",
@@ -520,11 +539,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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 pj = pretype_type empty_valcon env_p evdref lvar p in
+ let ccl = nf_evar (evars_of !evdref) 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;
+ let jtyp = inh_conv_coerce_to_tycon loc env evdref {uj_val = pred;
uj_type = typ} tycon
in
jtyp.uj_val, jtyp.uj_type
@@ -532,11 +551,11 @@ module Pretyping_F (Coercion : Coercion.S) = struct
let p = match tycon with
| Some (None, ty) -> ty
| None | Some _ ->
- e_new_evar isevars env ~src:(loc,InternalHole) (new_Type ())
+ e_new_evar evdref 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
+ let pred = nf_evar (evars_of !evdref) pred in
+ let p = nf_evar (evars_of !evdref) 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
@@ -555,118 +574,129 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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
+ let bj = pretype (mk_tycon pi) env_c evdref 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 IfStyle mis in
+ let ci = make_case_info env mis IfStyle 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)
+
+ | RCases (loc,sty,po,tml,eqns) ->
+ Cases.compile_cases loc sty
+ ((fun vtyc env evdref -> pretype vtyc env evdref lvar),evdref)
tycon env (* loc *) (po,tml,eqns)
| RCast (loc,c,k) ->
let cj =
match k with
CastCoerce ->
- let cj = pretype empty_tycon env isevars lvar c in
- evd_comb1 (Coercion.inh_coerce_to_base loc env) isevars cj
+ let cj = pretype empty_tycon env evdref lvar c in
+ evd_comb1 (Coercion.inh_coerce_to_base loc env) evdref cj
| CastConv (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
+ let tj = pretype_type empty_valcon env evdref lvar t in
+ let cj = pretype (mk_tycon tj.utj_val) env evdref 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
{ uj_val = v; uj_type = tj.utj_val }
in
- inh_conv_coerce_to_tycon loc env isevars cj tycon
+ inh_conv_coerce_to_tycon loc env evdref 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 j = (Retyping.get_judgment_of env (evars_of !evdref) c) in
j
- (*inh_conv_coerce_to_tycon loc env isevars j tycon*)
+ (*inh_conv_coerce_to_tycon loc env evdref 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
+ (* [pretype_type valcon env evdref lvar c] coerces [c] into a type *)
+ and pretype_type valcon env evdref lvar = function
| RHole loc ->
(match valcon with
| Some v ->
let s =
- let sigma = evars_of !isevars in
+ let sigma = evars_of !evdref 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
+ | Evar ev when is_Type (existential_type sigma ev) ->
+ evd_comb1 (define_evar_as_sort) evdref ev
| _ -> 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_val = e_new_evar evdref env ~src:loc (mkSort s);
utj_type = s})
| c ->
- let j = pretype empty_tycon env isevars lvar c in
+ let j = pretype empty_tycon env evdref lvar c in
let loc = loc_of_rawconstr c in
- let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) isevars j in
+ let tj = evd_comb1 (Coercion.inh_coerce_to_sort loc env) evdref j in
match valcon with
| None -> tj
| Some v ->
- if e_cumul env isevars v tj.utj_val then tj
+ if e_cumul env evdref v tj.utj_val then tj
else
error_unexpected_type_loc
- (loc_of_rawconstr c) env (evars_of !isevars) tj.utj_val v
+ (loc_of_rawconstr c) env (evars_of !evdref) tj.utj_val v
- let pretype_gen isevars env lvar kind c =
+ let pretype_gen_aux evdref 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
+ (pretype tycon env evdref lvar c).uj_val
| IsType ->
- (pretype_type empty_valcon env isevars lvar c).utj_val in
- nf_evar (evars_of !isevars) c'
-
+ (pretype_type empty_valcon env evdref lvar c).utj_val in
+ let evd,_ = consider_remaining_unif_problems env !evdref in
+ evdref := evd; c'
+
+ let pretype_gen evdref env lvar kind c =
+ let c = pretype_gen_aux evdref env lvar kind c in
+ evdref := Typeclasses.resolve_typeclasses ~onlyargs:true ~fail:false env !evdref;
+ nf_evar (evars_of !evdref) 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...
*)
let understand_judgment sigma env c =
- let isevars = ref (create_evar_defs sigma) in
- let j = pretype empty_tycon env isevars ([],[]) c in
- let isevars,_ = consider_remaining_unif_problems env !isevars 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
-
- let understand_judgment_tcc isevars env c =
- let j = pretype empty_tycon env isevars ([],[]) c in
- let sigma = evars_of !isevars in
+ let evdref = ref (create_evar_defs sigma) in
+ let j = pretype empty_tycon env evdref ([],[]) c in
+ let evd,_ = consider_remaining_unif_problems env !evdref in
+ let j = j_nf_evar (evars_of evd) j in
+ let evd = Typeclasses.resolve_typeclasses ~onlyargs:true ~fail:true env evd in
+ let j = j_nf_evar (evars_of evd) j in
+ check_evars env sigma evd (mkCast(j.uj_val,DEFAULTcast, j.uj_type));
+ j
+
+ let understand_judgment_tcc evdref env c =
+ let j = pretype empty_tycon env evdref ([],[]) c in
+ let sigma = evars_of !evdref in
let j = j_nf_evar sigma j in
- j
+ j
(* 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 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
- let isevars,_ = consider_remaining_unif_problems env !isevars in
- let c = nf_evar (evars_of isevars) c in
- if fail_evar then check_evars env sigma isevars c;
- isevars, c
+ let evdref = ref (Evd.create_evar_defs sigma) in
+ let c = pretype_gen evdref env lvar kind c in
+ let evd,_ = consider_remaining_unif_problems env !evdref in
+ if fail_evar then
+ let evd = Typeclasses.resolve_typeclasses ~onlyargs:false ~fail:true env evd in
+ let c = Evarutil.nf_isevar evd c in
+ check_evars env Evd.empty evd c;
+ evd, c
+ else evd, c
(** Entry points of the high-level type synthesis algorithm *)
@@ -682,12 +712,19 @@ module Pretyping_F (Coercion : Coercion.S) = struct
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 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
+ let understand_tcc_evars evdref env kind c =
+ pretype_gen evdref env ([],[]) kind c
+
+ let understand_tcc ?(resolve_classes=true) sigma env ?expected_type:exptyp c =
+ let evd, t =
+ if resolve_classes then
+ ise_pretype_gen false sigma env ([],[]) (OfType exptyp) c
+ else
+ let evdref = ref (Evd.create_evar_defs sigma) in
+ let c = pretype_gen_aux evdref env ([],[]) (OfType exptyp) c in
+ !evdref, nf_isevar !evdref c
+ in
+ Evd.evars_of evd, t
end
module Default : S = Pretyping_F(Coercion.Default)