diff options
author | herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2000-07-24 13:39:23 +0000 |
---|---|---|
committer | herbelin <herbelin@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2000-07-24 13:39:23 +0000 |
commit | 3afaf3dde673d77cacaabc354f008dfbe49a7cee (patch) | |
tree | 4264164faf763ab8d18274cd5aeffe5a1de21728 | |
parent | 83f038e61a4424fcf71aada9f97c91165b73aef8 (diff) |
Passage à des contextes de vars et de rels pouvant contenir des déclarations
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@568 85f007b7-540e-0410-9357-904b9bb8a0f7
79 files changed, 1938 insertions, 1739 deletions
diff --git a/contrib/omega/coq_omega.ml b/contrib/omega/coq_omega.ml index dd9b4b1e3..4375ef2a4 100644 --- a/contrib/omega/coq_omega.ml +++ b/contrib/omega/coq_omega.ml @@ -1352,7 +1352,7 @@ let normalize_equation id flag theorem pos t t1 t2 (tactic,defs) = else (tactic,defs) -let destructure_omega gl tac_def id c = +let destructure_omega gl tac_def (id,c) = if atompart_of_id id = "State" then tac_def else @@ -1388,7 +1388,7 @@ let destructure_omega gl tac_def id c = let coq_omega gl = clear_tables (); let tactic_normalisation, system = - it_sign (destructure_omega gl) ([],[]) (pf_untyped_hyps gl) in + List.fold_left (destructure_omega gl) ([],[]) (pf_hyps_types gl) in let prelude,sys = List.fold_left (fun (tac,sys) (t,(v,th,b)) -> @@ -1442,7 +1442,6 @@ let coq_omega = solver_time coq_omega let nat_inject gl = let aux = id_of_string "auxiliary" in let table = Hashtbl.create 7 in - let hyps = pf_untyped_hyps gl in let rec explore p t = try match destructurate t with | Kapp("plus",[t1;t2]) -> @@ -1472,14 +1471,14 @@ let nat_inject gl = (clever_rewrite_gen p (mk_minus (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_minus1),[t1;t2;VAR id])) - (loop [id] [mkAppL [| Lazy.force coq_le; t2;t1 |]])) + (loop [id,mkAppL [| Lazy.force coq_le; t2;t1 |]])) (explore (P_APP 1 :: p) t1)) (explore (P_APP 2 :: p) t2)); (tclTHEN (clever_rewrite_gen p (mk_integer 0) ((Lazy.force coq_inj_minus2),[t1;t2;VAR id])) - (loop [id] [mkAppL [| Lazy.force coq_gt; t2;t1 |]])) + (loop [id,mkAppL [| Lazy.force coq_gt; t2;t1 |]])) ] | Kapp("S",[t']) -> let rec is_number t = @@ -1513,10 +1512,9 @@ let nat_inject gl = | _ -> tclIDTAC with e when catchable_exception e -> tclIDTAC - and loop p0 p1 = - match (p0,p1) with - | ([], []) -> tclIDTAC - | ((i::li),(t::lt)) -> + and loop = function + | [] -> tclIDTAC + | (i,t)::lit -> begin try match destructurate t with Kapp("le",[t1;t2]) -> (tclTHEN @@ -1531,7 +1529,7 @@ let nat_inject gl = (explore [P_APP 2; P_TYPE] t2)) (clear [i])) (intros_using [i])) - (loop li lt)) + (loop lit)) | Kapp("lt",[t1;t2]) -> (tclTHEN (tclTHEN @@ -1545,7 +1543,7 @@ let nat_inject gl = (explore [P_APP 2; P_TYPE] t2)) (clear [i])) (intros_using [i])) - (loop li lt)) + (loop lit)) | Kapp("ge",[t1;t2]) -> (tclTHEN (tclTHEN @@ -1559,7 +1557,7 @@ let nat_inject gl = (explore [P_APP 2; P_TYPE] t2)) (clear [i])) (intros_using [i])) - (loop li lt)) + (loop lit)) | Kapp("gt",[t1;t2]) -> (tclTHEN (tclTHEN @@ -1573,7 +1571,7 @@ let nat_inject gl = (explore [P_APP 2; P_TYPE] t2)) (clear [i])) (intros_using [i])) - (loop li lt)) + (loop lit)) | Kapp("neq",[t1;t2]) -> (tclTHEN (tclTHEN @@ -1587,7 +1585,7 @@ let nat_inject gl = (explore [P_APP 2; P_TYPE] t2)) (clear [i])) (intros_using [i])) - (loop li lt)) + (loop lit)) | Kapp("eq",[typ;t1;t2]) -> if pf_conv_x gl typ (Lazy.force coq_nat) then (tclTHEN @@ -1602,13 +1600,12 @@ let nat_inject gl = (explore [P_APP 3; P_TYPE] t2)) (clear [i])) (intros_using [i])) - (loop li lt)) - else loop li lt - | _ -> loop li lt - with e when catchable_exception e -> loop li lt end - | (_, _) -> failwith "nat_inject" + (loop lit)) + else loop lit + | _ -> loop lit + with e when catchable_exception e -> loop lit end in - loop (List.rev (ids_of_sign hyps)) (List.rev (vals_of_sign hyps)) gl + loop (List.rev (pf_hyps_types gl)) gl let rec decidability gl t = match destructurate t with @@ -1649,16 +1646,15 @@ let rec decidability gl t = | _ -> error "Omega: Unrecognized proposition" let destructure_hyps gl = - let hyp = pf_untyped_hyps gl in - let rec loop evbd p0 p1 = match (p0,p1) with - | ([], []) -> (tclTHEN nat_inject coq_omega) - | ((i::li), (t::lt)) -> + let rec loop evbd = function + | [] -> (tclTHEN nat_inject coq_omega) + | (i,t)::lit -> begin try match destructurate t with - | Kapp(("Zle"|"Zge"|"Zgt"|"Zlt"|"Zne"),[t1;t2]) -> loop evbd li lt + | Kapp(("Zle"|"Zge"|"Zgt"|"Zlt"|"Zne"),[t1;t2]) -> loop evbd lit | Kapp("or",[t1;t2]) -> (tclTHENS ((tclTHEN ((tclTHEN (elim_id i) (clear [i]))) (intros_using [i]))) - ([ loop evbd (i::li) (t1::lt); loop evbd (i::li) (t2::lt) ])) + ([ loop evbd ((i,t1)::lit); loop evbd ((i,t2)::lit) ])) | Kapp("and",[t1;t2]) -> let i1 = id_of_string (string_of_id i ^ "_left") in let i2 = id_of_string (string_of_id i ^ "_right") in @@ -1666,7 +1662,7 @@ let destructure_hyps gl = (tclTHEN (tclTHEN (elim_id i) (clear [i])) (intros_using [i1;i2])) - (loop (i1::i2::evbd) (i1::i2::li) (t1::t2::lt))) + (loop (i1::i2::evbd) ((i1,t1)::(i2,t2)::lit))) | Kimp(t1,t2) -> if isprop (pf_type_of gl t1) & closed0 t2 then begin (tclTHEN @@ -1678,8 +1674,8 @@ let destructure_hyps gl = VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd (i::li) (mk_or (mk_not t1) t2 :: lt))) - end else loop evbd li lt + (loop evbd ((i,mk_or (mk_not t1) t2)::lit))) + end else loop evbd lit | Kapp("not",[t]) -> begin match destructurate t with Kapp("or",[t1;t2]) -> @@ -1690,8 +1686,7 @@ let destructure_hyps gl = t1; t2; VAR i |]]) (clear [i])) (intros_using [i])) - (loop evbd (i::li) - (mk_and (mk_not t1) (mk_not t2) :: lt))) + (loop evbd ((i,mk_and (mk_not t1) (mk_not t2)):: lit))) | Kapp("and",[t1;t2]) -> (tclTHEN (tclTHEN @@ -1701,8 +1696,7 @@ let destructure_hyps gl = decidability gl t1;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd (i::li) - (mk_or (mk_not t1) (mk_not t2) :: lt))) + (loop evbd ((i,mk_or (mk_not t1) (mk_not t2))::lit))) | Kimp(t1,t2) -> (tclTHEN (tclTHEN @@ -1712,7 +1706,7 @@ let destructure_hyps gl = decidability gl t1;VAR i |]]) (clear [i])) (intros_using [i])) - (loop evbd (i::li) (mk_and t1 (mk_not t2) :: lt))) + (loop evbd ((i,mk_and t1 (mk_not t2)) :: lit))) | Kapp("not",[t]) -> (tclTHEN (tclTHEN @@ -1722,7 +1716,7 @@ let destructure_hyps gl = decidability gl t; VAR i |]]) (clear [i])) (intros_using [i])) - (loop evbd (i::li) (t :: lt))) + (loop evbd ((i,t)::lit))) | Kapp("Zle", [t1;t2]) -> (tclTHEN (tclTHEN @@ -1731,7 +1725,7 @@ let destructure_hyps gl = t1;t2;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("Zge", [t1;t2]) -> (tclTHEN (tclTHEN @@ -1740,7 +1734,7 @@ let destructure_hyps gl = t1;t2;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("Zlt", [t1;t2]) -> (tclTHEN (tclTHEN @@ -1749,7 +1743,7 @@ let destructure_hyps gl = t1;t2;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("Zgt", [t1;t2]) -> (tclTHEN (tclTHEN @@ -1758,7 +1752,7 @@ let destructure_hyps gl = t1;t2;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("le", [t1;t2]) -> (tclTHEN (tclTHEN @@ -1767,7 +1761,7 @@ let destructure_hyps gl = t1;t2;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("ge", [t1;t2]) -> (tclTHEN (tclTHEN @@ -1776,7 +1770,7 @@ let destructure_hyps gl = t1;t2;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("lt", [t1;t2]) -> (tclTHEN (tclTHEN @@ -1785,7 +1779,7 @@ let destructure_hyps gl = t1;t2;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("gt", [t1;t2]) -> (tclTHEN (tclTHEN @@ -1794,7 +1788,7 @@ let destructure_hyps gl = t1;t2;VAR i|]]) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("eq",[typ;t1;t2]) -> if !old_style_flag then begin match destructurate (pf_nf gl typ) with @@ -1808,7 +1802,7 @@ let destructure_hyps gl = [t1;t2;VAR i]))) (clear [i])) (intros_using [i])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("Z",_) -> (tclTHEN (tclTHEN @@ -1819,30 +1813,29 @@ let destructure_hyps gl = [t1;t2;VAR i]))) (clear [i])) (intros_using [i])) - (loop evbd li lt)) - | _ -> loop evbd li lt + (loop evbd lit)) + | _ -> loop evbd lit end else begin match destructurate (pf_nf gl typ) with | Kapp("nat",_) -> (tclTHEN (convert_hyp i (mkAppL [| Lazy.force coq_neq; t1;t2|])) - (loop evbd li lt)) + (loop evbd lit)) | Kapp("Z",_) -> (tclTHEN (convert_hyp i (mkAppL [| Lazy.force coq_Zne; t1;t2|])) - (loop evbd li lt)) - | _ -> loop evbd li lt + (loop evbd lit)) + | _ -> loop evbd lit end - | _ -> loop evbd li lt + | _ -> loop evbd lit end - | _ -> loop evbd li lt - with e when catchable_exception e -> loop evbd li lt + | _ -> loop evbd lit + with e when catchable_exception e -> loop evbd lit end - | (_, _) -> anomaly "destructurate_hyps" in - loop (ids_of_sign hyp) (ids_of_sign hyp) (vals_of_sign hyp) gl + loop (pf_ids_of_hyps gl) (pf_hyps_types gl) gl let destructure_goal gl = let concl = pf_concl gl in diff --git a/kernel/declarations.ml b/kernel/declarations.ml index 510e1f12b..e83e1e509 100644 --- a/kernel/declarations.ml +++ b/kernel/declarations.ml @@ -19,7 +19,7 @@ type constant_body = { const_kind : path_kind; const_body : constant_value option; const_type : typed_type; - const_hyps : typed_type signature; + const_hyps : var_context; const_constraints : constraints; mutable const_opaque : bool } @@ -61,7 +61,7 @@ type one_inductive_body = { type mutual_inductive_body = { mind_kind : path_kind; mind_ntypes : int; - mind_hyps : typed_type signature; + mind_hyps : var_context; mind_packets : one_inductive_body array; mind_constraints : constraints; mind_singl : constr option; diff --git a/kernel/declarations.mli b/kernel/declarations.mli index 7853a2ce1..1dc4a7cb8 100644 --- a/kernel/declarations.mli +++ b/kernel/declarations.mli @@ -20,7 +20,7 @@ type constant_body = { const_kind : path_kind; const_body : constant_value option; const_type : typed_type; - const_hyps : typed_type signature; + const_hyps : var_context; (* New: younger hyp at top *) const_constraints : constraints; mutable const_opaque : bool } @@ -66,7 +66,7 @@ type one_inductive_body = { type mutual_inductive_body = { mind_kind : path_kind; mind_ntypes : int; - mind_hyps : typed_type signature; + mind_hyps : var_context; mind_packets : one_inductive_body array; mind_constraints : constraints; mind_singl : constr option; diff --git a/kernel/environ.ml b/kernel/environ.ml index 73a8c6487..217a7f989 100644 --- a/kernel/environ.ml +++ b/kernel/environ.ml @@ -26,13 +26,21 @@ type globals = { env_locals : (global * section_path) list; env_imports : import list } +type context = { + env_var_context : var_context; + env_rel_context : rel_context } + type env = { env_context : context; env_globals : globals; env_universes : universes } +let empty_context = { + env_var_context = empty_var_context; + env_rel_context = empty_rel_context } + let empty_env = { - env_context = ENVIRON (nil_sign,nil_dbsign); + env_context = empty_context; env_globals = { env_constants = Spmap.empty; env_inductives = Spmap.empty; @@ -43,29 +51,93 @@ let empty_env = { let universes env = env.env_universes let context env = env.env_context -let var_context env = let (ENVIRON(g,_)) = env.env_context in g +let var_context env = env.env_context.env_var_context +let rel_context env = env.env_context.env_rel_context (* Construction functions. *) -let push_var idvar env = - { env with env_context = add_glob idvar env.env_context } - -let change_hyps f env = - let (ENVIRON(g,r)) = env.env_context in - { env with env_context = ENVIRON (f g, r) } - -(* == functions to deal with names in contexts (previously in cases.ml) == *) - -(* If cur=(Rel j) then - * if env = ENVIRON(sign,[na_h:Th]...[na_j:Tj]...[na_1:T1]) - * then it yields ENVIRON(sign,[na_h:Th]...[Name id:Tj]...[na_1:T1]) - *) -let change_name_rel env j id = - { env with env_context = change_name_env (context env) j id } -(****) - -let push_rel idrel env = - { env with env_context = add_rel idrel env.env_context } +let map_context f env = + let context = env.env_context in + { env with + env_context = { + context with + env_var_context = map_var_context f context.env_var_context ; + env_rel_context = map_rel_context f context.env_rel_context } } + +let var_context_app f env = + { env with + env_context = { env.env_context with + env_var_context = f env.env_context.env_var_context } } + +let change_hyps = var_context_app + +let push_var d = var_context_app (add_var d) +let push_var_def def = var_context_app (add_var_def def) +let push_var_decl decl = var_context_app (add_var_decl decl) +let pop_var id = var_context_app (pop_var id) + +let rel_context_app f env = + { env with + env_context = { env.env_context with + env_rel_context = f env.env_context.env_rel_context } } + +let reset_context env = + { env with + env_context = { env_var_context = empty_var_context; + env_rel_context = empty_rel_context} } + +let fold_var_context f env a = + snd (Sign.fold_var_context + (fun d (env,e) -> (push_var d env, f env d e)) + (var_context env) (reset_context env,a)) + +let process_var_context f env = + Sign.fold_var_context + (fun d env -> f env d) (var_context env) (reset_context env) + +let process_var_context_both_sides f env = + fold_var_context_both_sides f (var_context env) (reset_context env) + +let push_rel d = rel_context_app (add_rel d) +let push_rel_def def = rel_context_app (add_rel_def def) +let push_rel_decl decl = rel_context_app (add_rel_decl decl) +let push_rels ctxt = rel_context_app (concat_rel_context ctxt) + +let push_rels_to_vars env = + let sign0 = env.env_context.env_var_context in + let (subst,_,sign) = + List.fold_right + (fun (na,c,t) (subst,avoid,sign) -> + let na = if na = Anonymous then Name(id_of_string"_") else na in + let id = next_name_away na avoid in + ((VAR id)::subst,id::avoid, + add_var (id,option_app (substl subst) c,typed_app (substl subst) t) + sign)) + env.env_context.env_rel_context ([],ids_of_var_context sign0,sign0) + in subst, (var_context_app (fun _ -> sign) env) + +let reset_rel_context env = + { env with + env_context = { env_var_context = env.env_context.env_var_context; + env_rel_context = empty_rel_context} } + +let fold_rel_context f env a = + snd (List.fold_right + (fun d (env,e) -> (push_rel d env, f env d e)) + (rel_context env) (reset_rel_context env,a)) + +let process_rel_context f env = + List.fold_right (fun d env -> f env d) + (rel_context env) (reset_rel_context env) + +let instantiate_vars = instantiate_sign + +let ids_of_context env = + (ids_of_rel_context env.env_context.env_rel_context) + @ (ids_of_var_context env.env_context.env_var_context) + +let names_of_rel_context env = + names_of_rel_context env.env_context.env_rel_context let set_universes g env = if env.env_universes == g then env else { env with env_universes = g } @@ -109,12 +181,19 @@ let new_meta ()=incr meta_ctr;!meta_ctr;; (* Access functions. *) -let lookup_var id env = - let (_,var) = lookup_glob id env.env_context in - (Name id, var) +let lookup_var_type id env = + lookup_id_type id env.env_context.env_var_context + +let lookup_var_value id env = + lookup_id_value id env.env_context.env_var_context + +let lookup_var id env = lookup_id id env.env_context.env_var_context -let lookup_rel n env = - Sign.lookup_rel n env.env_context +let lookup_rel_type n env = + Sign.lookup_rel_type n env.env_context.env_rel_context + +let lookup_rel_value n env = + Sign.lookup_rel_value n env.env_context.env_rel_context let lookup_constant sp env = Spmap.find sp env.env_globals.env_constants @@ -171,7 +250,7 @@ let hdchar env c = | Rel n -> (if n<=k then "p" (* the initial term is flexible product/function *) else - try match lookup_rel (n-k) env with + try match lookup_rel_type (n-k) env with | Name id,_ -> lowercase_first_char id | Anonymous,t -> hdrec 0 (lift (n-k) (body_of_type t)) with Not_found -> "y") @@ -196,6 +275,32 @@ let it_lambda_name env = List.fold_left (fun c (n,t) ->lambda_name env (n,t,c)) let prod_create env (a,b) = mkProd (named_hd env a Anonymous) a b let lambda_create env (a,b) = mkLambda (named_hd env a Anonymous) a b +let name_assumption env (na,c,t) = + match c with + | None -> (named_hd env (body_of_type t) na, None, t) + | Some body -> (named_hd env body na, c, t) + +let prod_assum_name env b d = mkProd_or_LetIn (name_assumption env d) b +let lambda_assum_name env b d = mkLambda_or_LetIn (name_assumption env d) b + +let it_mkProd_or_LetIn_name env = List.fold_left (prod_assum_name env) +let it_mkLambda_or_LetIn_name env = List.fold_left (lambda_assum_name env) + +let it_mkProd_or_LetIn = List.fold_left (fun c d -> mkProd_or_LetIn d c) +let it_mkLambda_or_LetIn = List.fold_left (fun c d -> mkLambda_or_LetIn d c) + +let it_mkNamedProd_or_LetIn = it_var_context_quantifier mkNamedProd_or_LetIn +let it_mkNamedLambda_or_LetIn = it_var_context_quantifier mkNamedLambda_or_LetIn + +let make_all_name_different env = + let avoid = ref (ids_of_var_context (var_context env)) in + process_rel_context + (fun newenv (na,c,t) -> + let id = next_name_away na !avoid in + avoid := id::!avoid; + push_rel (Name id,c,t) newenv) + env + (* Abstractions. *) let evaluable_abst env = function diff --git a/kernel/environ.mli b/kernel/environ.mli index c8e3642be..616a11162 100644 --- a/kernel/environ.mli +++ b/kernel/environ.mli @@ -10,26 +10,64 @@ open Univ open Sign (*i*) -(* Unsafe environments. We define here a datatype for environments. +(*s Unsafe environments. We define here a datatype for environments. Since typing is not yet defined, it is not possible to check the informations added in environments, and that is what we speak here of ``unsafe'' environments. *) +type context type env +val empty_context : context val empty_env : env val universes : env -> universes val context : env -> context +val rel_context : env -> rel_context val var_context : env -> var_context -val push_var : identifier * typed_type -> env -> env -val change_hyps : - (typed_type signature -> typed_type signature) -> env -> env -val change_name_rel : env -> int -> identifier -> env - -val push_rel : name * typed_type -> env -> env - +(* This forgets var and rel contexts *) +val reset_context : env -> env + +(*s This concerns only the context of local vars referred by names + [var_context] *) +val push_var : var_declaration -> env -> env +val push_var_decl : identifier * typed_type -> env -> env +val push_var_def : identifier * constr * typed_type -> env -> env +val change_hyps : (var_context -> var_context) -> env -> env +val instantiate_vars : var_context -> constr list -> (identifier * constr) list +val pop_var : identifier -> env -> env + +(*s This concerns only the context of local vars referred by indice + [rel_context] *) +val push_rel : rel_declaration -> env -> env +val push_rel_decl : name * typed_type -> env -> env +val push_rel_def : name * constr * typed_type -> env -> env +val push_rels : rel_context -> env -> env +val names_of_rel_context : env -> names_context + +(*s Returns also the substitution to be applied to rel's *) +val push_rels_to_vars : env -> constr list * env + +(* Gives identifiers in var_context and rel_context *) +val ids_of_context : env -> identifier list +val map_context : (constr -> constr) -> env -> env + +(*s Recurrence on var_context *) +val fold_var_context : (env -> var_declaration -> 'a -> 'a) -> env -> 'a -> 'a +val process_var_context : (env -> var_declaration -> env) -> env -> env + +(* [process_var_context_both_sides f env] iters [f] on the var + declarations of env taking as argument both the initial segment, the + middle value and the tail segment *) +val process_var_context_both_sides : + (env -> var_declaration -> var_context -> env) -> env -> env + +(*s Recurrence on rel_context *) +val fold_rel_context : (env -> rel_declaration -> 'a -> 'a) -> env -> 'a -> 'a +val process_rel_context : (env -> rel_declaration -> env) -> env -> env + +(*s add entries to environment *) val set_universes : universes -> env -> env val add_constraints : constraints -> env -> env val add_constant : @@ -41,13 +79,34 @@ val add_abstraction : val new_meta : unit -> int -val lookup_var : identifier -> env -> name * typed_type -val lookup_rel : int -> env -> name * typed_type +(*s Looks up in environment *) + +(* Looks up in the context of local vars referred by names (var_context) *) +(* raises [Not_found] if the identifier is not found *) +val lookup_var_type : identifier -> env -> typed_type +val lookup_var_value : identifier -> env -> constr option +val lookup_var : identifier -> env -> constr option * typed_type + +(* Looks up in the context of local vars referred by indice (rel_context) *) +(* raises [Not_found] if the index points out of the context *) +val lookup_rel_type : int -> env -> name * typed_type +val lookup_rel_value : int -> env -> constr option + +(* Looks up in the context of global constant names *) +(* raises [Not_found] if the required path is not found *) val lookup_constant : section_path -> env -> constant_body + +(* Looks up in the context of global inductive names *) +(* raises [Not_found] if the required path is not found *) val lookup_mind : section_path -> env -> mutual_inductive_body +(*s Miscellanous *) val id_of_global : env -> sorts oper -> identifier +val make_all_name_different : env -> env + +(*s Functions creating names for anonymous names *) + val id_of_name_using_hdchar : env -> constr -> name -> identifier (* [named_hd env t na] just returns [na] is it defined, otherwise it @@ -68,6 +127,15 @@ val prod_name : env -> name * constr * constr -> constr val it_lambda_name : env -> constr -> (name * constr) list -> constr val it_prod_name : env -> constr -> (name * constr) list -> constr +val it_mkLambda_or_LetIn_name : env -> constr -> rel_context -> constr +val it_mkProd_or_LetIn_name : env -> constr -> rel_context -> constr + +val it_mkLambda_or_LetIn : constr -> rel_context -> constr +val it_mkProd_or_LetIn : constr -> rel_context -> constr + +val it_mkNamedLambda_or_LetIn : constr -> var_context -> constr +val it_mkNamedProd_or_LetIn : constr -> var_context -> constr + (* [lambda_create env (t,c)] builds [[x:t]c] where [x] is a name built from [t]; [prod_create env (t,c)] builds [(x:t)c] where [x] is a name built from [t] *) diff --git a/kernel/evd.mli b/kernel/evd.mli index e9f7818ff..3ae00f651 100644 --- a/kernel/evd.mli +++ b/kernel/evd.mli @@ -48,6 +48,6 @@ val is_evar : 'a evar_map -> evar -> bool val is_defined : 'a evar_map -> evar -> bool -val evar_hyps : 'a evar_info -> typed_type signature +val evar_hyps : 'a evar_info -> var_context val id_of_existential : evar -> identifier diff --git a/kernel/generic.ml b/kernel/generic.ml index 70acca45d..e7e1acc2f 100644 --- a/kernel/generic.ml +++ b/kernel/generic.ml @@ -304,10 +304,14 @@ let replace_vars var_alist = in if var_alist = [] then (function x -> x) else substrec 0 -let subst_varn str n = replace_vars [(str, (Rel n))] - (* (subst_var str t) substitute (VAR str) by (Rel 1) in t *) -let subst_var str = subst_varn str 1 +let subst_var str = replace_vars [(str, Rel 1)] + +(* (subst_vars [id1;...;idn] t) substitute (VAR idj) by (Rel j) in t *) +let subst_vars vars = + let _,subst = + List.fold_left (fun (n,l) var -> ((n+1),(var,Rel n)::l)) (1,[]) vars + in replace_vars (List.rev subst) (* [Rel (n+m);...;Rel(n+1)] *) diff --git a/kernel/generic.mli b/kernel/generic.mli index 3f14efaee..8c2991ff6 100644 --- a/kernel/generic.mli +++ b/kernel/generic.mli @@ -60,8 +60,12 @@ val subst1 : 'a term -> 'a term -> 'a term val global_vars : 'a term -> identifier list val global_vars_set : 'a term -> Idset.t -val subst_var : identifier -> 'a term -> 'a term val replace_vars : (identifier * 'a term) list -> 'a term -> 'a term +val subst_var : identifier -> 'a term -> 'a term + +(* [subst_vars [id1;...;idn] t] substitute [VAR idj] by [Rel j] in [t] + if two names are identical, the one of least indice is keeped *) +val subst_vars : identifier list -> 'a term -> 'a term (* [rel_list n m] and [rel_vect n m] compute [[Rel (n+m);...;Rel(n+1)]] *) val rel_vect : int -> int -> 'a term array diff --git a/kernel/indtypes.ml b/kernel/indtypes.ml index 197f5c299..0536b1f2f 100644 --- a/kernel/indtypes.ml +++ b/kernel/indtypes.ml @@ -27,10 +27,10 @@ open Typeops type inductive_error = (* These are errors related to inductive constructions in this module *) - | NonPos of name list * constr * constr - | NotEnoughArgs of name list * constr * constr - | NotConstructor of name list * constr * constr - | NonPar of name list * constr * int * constr * constr + | NonPos of env * constr * constr + | NotEnoughArgs of env * constr * constr + | NotConstructor of env * constr * constr + | NonPar of env * constr * int * constr * constr | SameNamesTypes of identifier | SameNamesConstructors of identifier * identifier | NotAnArity of identifier @@ -74,30 +74,22 @@ let mind_check_names mie = for all the given types. *) let mind_extract_params n c = - let (l,c') = decompose_prod_n n c in (List.rev l,c') + let (l,c') = decompose_prod_n_assum n c in (l,c') let extract nparams c = - try mind_extract_params nparams c + try fst (mind_extract_params nparams c) with UserError _ -> raise (InductiveError BadEntry) -let check_params nparams params c = - let eparams = fst (extract nparams c) in - try - List.iter2 - (fun (n1,t1) (n2,t2) -> - if n1 <> n2 || strip_all_casts t1 <> strip_all_casts t2 then - raise (InductiveError BadEntry)) - eparams params - with Invalid_argument _ -> - raise (InductiveError BadEntry) +let check_params params params' = + if not (params = params') then raise (InductiveError BadEntry) let mind_extract_and_check_params mie = let nparams = mie.mind_entry_nparams in match mie.mind_entry_inds with | [] -> anomaly "empty inductive types declaration" | (_,ar,_,_)::l -> - let (params,_) = extract nparams ar in - List.iter (fun (_,c,_,_) -> check_params nparams params c) l; + let params = extract nparams ar in + List.iter (fun (_,c,_,_) -> check_params params (extract nparams c)) l; params let decomp_all_DLAMV_name constr = @@ -110,7 +102,8 @@ let decomp_all_DLAMV_name constr = let mind_check_lc params mie = let nparams = List.length params in - let check_lc (_,_,_,lc) = List.iter (check_params nparams params) lc in + let check_lc (_,_,_,lc) = + List.iter (fun c -> check_params params (extract nparams c)) lc in List.iter check_lc mie.mind_entry_inds let mind_check_arities env mie = @@ -130,8 +123,6 @@ let allowed_sorts issmall isunit = function | Prop Null -> if isunit then [prop;spec] else [prop] -let is_small_type t = is_small t.typ - type ill_formed_ind = | LocalNonPos of int | LocalNotEnoughArgs of int @@ -140,9 +131,9 @@ type ill_formed_ind = exception IllFormedInd of ill_formed_ind -let explain_ind_err ntyp lna nbpar c err = +let explain_ind_err ntyp env0 nbpar c err = let (lpar,c') = mind_extract_params nbpar c in - let env = (List.map fst lpar)@lna in + let env = push_rels lpar env0 in match err with | LocalNonPos kt -> raise (InductiveError (NonPos (env,c',Rel (kt+nbpar)))) @@ -290,14 +281,13 @@ let listrec_mconstr env ntypes nparams i indlc = List.rev lrec in check_constr_rec [] in - let lna = it_dbenv (fun l na t -> na::l) [] (context env) in Array.map (fun c -> let c = body_of_type c in try check_construct true (1+nparams) (decomp_par nparams c) with IllFormedInd err -> - explain_ind_err (ntypes-i+1) lna nparams c err) + explain_ind_err (ntypes-i+1) env nparams c err) indlc let is_recursive listind = @@ -357,7 +347,7 @@ let cci_inductive env env_ar kind nparams finite inds cst = let packets = Array.of_list (list_map_i one_packet 1 inds) in { mind_kind = kind; mind_ntypes = ntypes; - mind_hyps = keep_hyps (var_context env) ids; + mind_hyps = keep_hyps ids (var_context env); mind_packets = packets; mind_constraints = cst; mind_singl = None; diff --git a/kernel/indtypes.mli b/kernel/indtypes.mli index 5a7f040d9..46c0255a2 100644 --- a/kernel/indtypes.mli +++ b/kernel/indtypes.mli @@ -15,10 +15,10 @@ open Environ type inductive_error = (* These are errors related to inductive constructions in this module *) - | NonPos of name list * constr * constr - | NotEnoughArgs of name list * constr * constr - | NotConstructor of name list * constr * constr - | NonPar of name list * constr * int * constr * constr + | NonPos of env * constr * constr + | NotEnoughArgs of env * constr * constr + | NotConstructor of env * constr * constr + | NonPar of env * constr * int * constr * constr | SameNamesTypes of identifier | SameNamesConstructors of identifier * identifier | NotAnArity of identifier @@ -45,11 +45,11 @@ val mind_check_names : mutual_inductive_entry -> unit [mind_entry_inds]. *) val mind_extract_and_check_params : - mutual_inductive_entry -> (name * constr) list + mutual_inductive_entry -> Sign.rel_context -val mind_extract_params : int -> constr -> (name * constr) list * constr +val mind_extract_params : int -> constr -> Sign.rel_context * constr -val mind_check_lc : (name * constr) list -> mutual_inductive_entry -> unit +val mind_check_lc : Sign.rel_context -> mutual_inductive_entry -> unit (* [mind_check_arities] checks that the types declared for all the inductive types are some arities. *) diff --git a/kernel/inductive.ml b/kernel/inductive.ml index cdfd27cd0..43b756651 100644 --- a/kernel/inductive.ml +++ b/kernel/inductive.ml @@ -18,6 +18,11 @@ type inductive_instance = { mis_args : constr array; mis_mip : one_inductive_body } + +let build_mis ((sp,tyi),args) mib = + { mis_sp = sp; mis_mib = mib; mis_tyi = tyi; mis_args = args; + mis_mip = mind_nth_type_packet mib tyi } + let mis_ntypes mis = mis.mis_mib.mind_ntypes let mis_nparams mis = mis.mis_mib.mind_nparams @@ -36,9 +41,9 @@ let mis_consnames mis = mis.mis_mip.mind_consnames let mis_inductive mis = ((mis.mis_sp,mis.mis_tyi),mis.mis_args) let mis_typed_lc mis = - let ids = ids_of_sign mis.mis_mib.mind_hyps in + let sign = mis.mis_mib.mind_hyps in let args = Array.to_list mis.mis_args in - Array.map (fun t -> Instantiate.instantiate_type ids t args) + Array.map (fun t -> Instantiate.instantiate_type sign t args) mis.mis_mip.mind_nf_lc let mis_lc mis = Array.map body_of_type (mis_typed_lc mis) @@ -75,9 +80,9 @@ let mis_type_mconstruct i mispec = typed_app (substl (list_tabulate make_Ik ntypes)) specif.(i-1) let mis_typed_arity mis = - let idhyps = ids_of_sign mis.mis_mib.mind_hyps + let hyps = mis.mis_mib.mind_hyps and largs = Array.to_list mis.mis_args in - Instantiate.instantiate_type idhyps mis.mis_mip.mind_nf_arity largs + Instantiate.instantiate_type hyps mis.mis_mip.mind_nf_arity largs (* let mis_arity mispec = incast_type (mis_typed_arity mispec) @@ -87,7 +92,7 @@ let mis_arity mis = body_of_type (mis_typed_arity mis) let mis_params_ctxt mispec = let paramsign,_ = - decompose_prod_n mispec.mis_mib.mind_nparams + decompose_prod_n_assum mispec.mis_mib.mind_nparams (body_of_type (mis_typed_arity mispec)) in paramsign @@ -203,10 +208,8 @@ let find_mcoinductype env sigma c = | _ -> raise Induc (* raise Induc if not an inductive type *) -let lookup_mind_specif ((sp,tyi),args) env = - let mib = lookup_mind sp env in - { mis_sp = sp; mis_mib = mib; mis_tyi = tyi; mis_args = args; - mis_mip = mind_nth_type_packet mib tyi } +let lookup_mind_specif ((sp,tyi),args as ind) env = + build_mis ind (lookup_mind sp env) let find_rectype env sigma ty = let (mind,largs) = find_mrectype env sigma ty in @@ -219,7 +222,7 @@ type constructor_summary = { cs_cstr : constructor; cs_params : constr list; cs_nargs : int; - cs_args : (name * constr) list; + cs_args : rel_context; cs_concl_realargs : constr array } @@ -227,18 +230,18 @@ let lift_constructor n cs = { cs_cstr = (let (csp,ctxt) = cs.cs_cstr in (csp,Array.map (lift n) ctxt)); cs_params = List.map (lift n) cs.cs_params; cs_nargs = cs.cs_nargs; - cs_args = lift_context n cs.cs_args; + cs_args = lift_rel_context n cs.cs_args; cs_concl_realargs = Array.map (liftn n (cs.cs_nargs+1)) cs.cs_concl_realargs } let get_constructor (IndFamily (mispec,params)) j = assert (j <= mis_nconstr mispec); let typi = mis_type_nf_mconstruct j mispec in - let (args,ccl) = decompose_prod (prod_applist typi params) in + let (args,ccl) = decompose_prod_assum (prod_applist typi params) in let (_,vargs) = array_chop (mis_nparams mispec + 1) (destAppL (ensure_appl ccl)) in { cs_cstr = ith_constructor_of_inductive (mis_inductive mispec) j; cs_params = params; - cs_nargs = List.length args; + cs_nargs = rel_context_length args; cs_args = args; cs_concl_realargs = vargs } @@ -277,8 +280,8 @@ let make_arity env dep indf s = let build_branch_type env dep p cs = let base = appvect (lift cs.cs_nargs p, cs.cs_concl_realargs) in if dep then - it_prod_name env + it_mkProd_or_LetIn_name env (applist (base,[build_dependent_constructor cs])) cs.cs_args else - prod_it base cs.cs_args + it_mkProd_or_LetIn base cs.cs_args diff --git a/kernel/inductive.mli b/kernel/inductive.mli index 97118a517..3b45dad0b 100644 --- a/kernel/inductive.mli +++ b/kernel/inductive.mli @@ -27,6 +27,8 @@ corresponds various appropriated functions *) type inductive_instance (* ex-[mind_specif] *) +val build_mis : inductive -> mutual_inductive_body -> inductive_instance + val mis_index : inductive_instance -> int val mis_ntypes : inductive_instance -> int val mis_nconstr : inductive_instance -> int @@ -43,7 +45,7 @@ val mis_consnames : inductive_instance -> identifier array val mis_typed_arity : inductive_instance -> typed_type val mis_inductive : inductive_instance -> inductive val mis_arity : inductive_instance -> constr -val mis_params_ctxt : inductive_instance -> (name * constr) list +val mis_params_ctxt : inductive_instance -> rel_context val mis_sort : inductive_instance -> sorts val mis_type_mconstruct : int -> inductive_instance -> typed_type @@ -98,7 +100,7 @@ type constructor_summary = { cs_cstr : constructor; cs_params : constr list; cs_nargs : int; - cs_args : (name * constr) list; + cs_args : rel_context; cs_concl_realargs : constr array } @@ -163,6 +165,8 @@ val find_rectype : env -> 'a evar_map -> constr -> inductive_type val get_constructors : inductive_family -> constructor_summary array +val get_constructor : inductive_family -> int -> constructor_summary + (* [get_arity inf] returns the product and the sort of the arity of the inductive family described by [is]; global parameters are already instanciated; [indf] must be defined w.r.t. [env] and [sigma]; diff --git a/kernel/instantiate.ml b/kernel/instantiate.ml index 5ec0e01b0..9fb85961f 100644 --- a/kernel/instantiate.ml +++ b/kernel/instantiate.ml @@ -11,21 +11,22 @@ open Evd open Declarations open Environ -let is_id_inst ids args = - let is_id id = function - | VAR id' -> id = id' +let is_id_inst inst = + let is_id = function + | id, VAR id' -> id = id' | _ -> false in - List.for_all2 is_id ids args + List.for_all is_id inst -let instantiate_constr ids c args = - if is_id_inst ids args then +let instantiate_constr sign c args = + let inst = instantiate_vars sign args in + if is_id_inst inst then c else - replace_vars (List.combine ids args) c + replace_vars inst c -let instantiate_type ids tty args = - typed_app (fun c -> instantiate_constr ids c args) tty +let instantiate_type sign tty args = + typed_app (fun c -> instantiate_constr sign c args) tty (* Constants. *) @@ -33,8 +34,7 @@ let instantiate_type ids tty args = let constant_type env sigma (sp,args) = let cb = lookup_constant sp env in (* TODO: check args *) - instantiate_type - (ids_of_sign cb.const_hyps) cb.const_type (Array.to_list args) + instantiate_type cb.const_hyps cb.const_type (Array.to_list args) type const_evaluation_error = NotDefinedConst | OpaqueConst @@ -48,8 +48,7 @@ let constant_value env cst = match cb.const_body with | Some v -> let body = cook_constant v in - instantiate_constr - (ids_of_sign cb.const_hyps) body (Array.to_list args) + instantiate_constr cb.const_hyps body (Array.to_list args) | None -> anomalylabstrm "termenv__constant_value" [< 'sTR "a defined constant with no body." >] @@ -62,14 +61,14 @@ let existential_type sigma (n,args) = let info = Evd.map sigma n in let hyps = evar_hyps info in (* TODO: check args [this comment was in Typeops] *) - instantiate_constr (ids_of_sign hyps) info.evar_concl (Array.to_list args) + instantiate_constr hyps info.evar_concl (Array.to_list args) let existential_value sigma (n,args) = let info = Evd.map sigma n in let hyps = evar_hyps info in match info.evar_body with | Evar_defined c -> - instantiate_constr (ids_of_sign hyps) c (Array.to_list args) + instantiate_constr hyps c (Array.to_list args) | Evar_empty -> anomaly "a defined existential with no body" diff --git a/kernel/instantiate.mli b/kernel/instantiate.mli index 40601c604..d3454f1cc 100644 --- a/kernel/instantiate.mli +++ b/kernel/instantiate.mli @@ -5,15 +5,16 @@ open Names open Term open Evd +open Sign open Environ (*i*) (* Instantiation of constants and inductives on their real arguments. *) val instantiate_constr : - identifier list -> constr -> constr list -> constr + var_context -> constr -> constr list -> constr val instantiate_type : - identifier list -> typed_type -> constr list -> typed_type + var_context -> typed_type -> constr list -> typed_type type const_evaluation_error = NotDefinedConst | OpaqueConst exception NotEvaluableConst of const_evaluation_error diff --git a/kernel/names.mli b/kernel/names.mli index 1ccf3c12c..84b25d9cd 100644 --- a/kernel/names.mli +++ b/kernel/names.mli @@ -5,7 +5,7 @@ open Pp (*i*) -(* Names. *) +(*s Identifiers *) type identifier type name = Name of identifier | Anonymous @@ -41,6 +41,7 @@ type path_kind = CCI | FW | OBJ val string_of_kind : path_kind -> string val kind_of_string : string -> path_kind +(*s Section paths *) type section_path val make_path : string list -> identifier -> path_kind -> section_path @@ -71,6 +72,7 @@ val sp_gt : section_path * section_path -> bool module Spset : Set.S with type elt = section_path module Spmap : Map.S with type key = section_path +(*s Specific paths for declarations *) type inductive_path = section_path * int type constructor_path = inductive_path * int diff --git a/kernel/reduction.ml b/kernel/reduction.ml index c7acd792a..c4642b933 100644 --- a/kernel/reduction.ml +++ b/kernel/reduction.ml @@ -530,32 +530,27 @@ let contract_fix ((recindices,bodynum),(types,names,bodies as typedbodies)) = let make_Fi j = mkFix ((recindices,nbodies-j-1),typedbodies) in substl (list_tabulate make_Fi nbodies) bodies.(bodynum) -let fix_recarg fix stack = - match fix with - | DOPN(Fix(recindices,bodynum),_) -> - if 0 <= bodynum & bodynum < Array.length recindices then - let recargnum = Array.get recindices bodynum in - (try - Some (recargnum, List.nth stack recargnum) - with Failure "nth" | Invalid_argument "List.nth" -> - None) - else - None - | _ -> assert false +let fix_recarg ((recindices,bodynum),_) stack = + if 0 <= bodynum & bodynum < Array.length recindices then + let recargnum = Array.get recindices bodynum in + (try + Some (recargnum, List.nth stack recargnum) + with Failure "nth" | Invalid_argument "List.nth" -> + None) + else + None let reduce_fix whfun fix stack = - match fix with - | DOPN(Fix(recindices,bodynum),bodyvect) -> - (match fix_recarg fix stack with - | None -> (false,(fix,stack)) - | Some (recargnum,recarg) -> - let (recarg'hd,_ as recarg') = whfun recarg [] in - let stack' = list_assign stack recargnum (applist recarg') in - (match recarg'hd with - | DOPN(MutConstruct _,_) -> - (true,(contract_fix (destFix fix),stack')) - | _ -> (false,(fix,stack')))) - | _ -> assert false + let dfix = destFix fix in + match fix_recarg dfix stack with + | None -> (false,(fix,stack)) + | Some (recargnum,recarg) -> + let (recarg'hd,_ as recarg') = whfun recarg [] in + let stack' = list_assign stack recargnum (applist recarg') in + (match recarg'hd with + | DOPN(MutConstruct _,_) -> + (true,(contract_fix dfix,stack')) + | _ -> (false,(fix,stack'))) (* NB : Cette fonction alloue peu c'est l'appel ``let (recarg'hd,_ as recarg') = whfun recarg [] in'' @@ -1035,10 +1030,11 @@ let sort_of_arity env c = snd (splay_arity env Evd.empty c) let decomp_n_prod env sigma n = let rec decrec m ln c = if m = 0 then (ln,c) else match whd_betadeltaiota env sigma c with - | DOP2(Prod,a,DLAM(n,c_0)) -> decrec (m-1) ((n,a)::ln) c_0 + | DOP2(Prod,a,DLAM(n,c0)) -> + decrec (m-1) (Sign.add_rel_decl (n,outcast_type a) ln) c0 | _ -> error "decomp_n_prod: Not enough products" in - decrec n [] + decrec n Sign.empty_rel_context diff --git a/kernel/reduction.mli b/kernel/reduction.mli index fb623595a..bab5f446f 100644 --- a/kernel/reduction.mli +++ b/kernel/reduction.mli @@ -97,7 +97,7 @@ val splay_prod : env -> 'a evar_map -> constr -> (name * constr) list * constr val splay_arity : env -> 'a evar_map -> constr -> (name * constr) list * sorts val sort_of_arity : env -> constr -> sorts val decomp_n_prod : - env -> 'a evar_map -> int -> constr -> ((name * constr) list) * constr + env -> 'a evar_map -> int -> constr -> Sign.rel_context * constr type 'a miota_args = { mP : constr; (* the result type *) @@ -130,7 +130,7 @@ val fold_commands : constr list -> 'a reduction_function val pattern_occs : (int list * constr * constr) list -> 'a reduction_function val compute : 'a reduction_function -val fix_recarg : constr -> 'a list -> (int * 'a) option +val fix_recarg : fixpoint -> 'a list -> (int * 'a) option val reduce_fix : (constr -> 'a list -> constr * constr list) -> constr -> constr list -> bool * (constr * constr list) diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index d4b610a8e..c99bd4bbb 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -48,7 +48,7 @@ let rec execute mf env cstr = (relative env Evd.empty n, cst0) | IsVar id -> - (make_judge cstr (snd (lookup_var id env)), cst0) + (make_judge cstr (lookup_var_type id env), cst0) | IsAbst _ -> if evaluable_abst env cstr then @@ -108,7 +108,7 @@ let rec execute mf env cstr = | IsLambda (name,c1,c2) -> let (j,cst1) = execute mf env c1 in let var = assumption_of_judgment env Evd.empty j in - let env1 = push_rel (name,var) env in + let env1 = push_rel_decl (name,var) env in let (j',cst2) = execute mf env1 c2 in let (j,cst3) = abs_rel env1 Evd.empty name var j' in let cst = Constraint.union cst1 (Constraint.union cst2 cst3) in @@ -117,7 +117,7 @@ let rec execute mf env cstr = | IsProd (name,c1,c2) -> let (j,cst1) = execute mf env c1 in let varj = type_judgment env Evd.empty j in - let env1 = push_rel (name,assumption_of_type_judgment varj) env in + let env1 = push_rel_decl (name,assumption_of_type_judgment varj) env in let (j',cst2) = execute mf env1 c2 in let (j,cst3) = gen_rel env1 Evd.empty name varj j' in let cst = Constraint.union cst1 (Constraint.union cst2 cst3) in @@ -137,7 +137,7 @@ and execute_fix mf env lar lfi vdef = let nlara = List.combine (List.rev lfi) (Array.to_list (vect_lift_type lara)) in let env1 = - List.fold_left (fun env nvar -> push_rel nvar env) env nlara in + List.fold_left (fun env nvar -> push_rel_decl nvar env) env nlara in let (vdefj,cst2) = execute_array mf env1 vdef in let vdefv = Array.map j_val_only vdefj in let cst3 = type_fixpoint env1 Evd.empty lfi lara vdefj in @@ -239,8 +239,9 @@ let universes = universes let context = context let var_context = var_context +let lookup_var_type = lookup_var_type +let lookup_rel_type = lookup_rel_type let lookup_var = lookup_var -let lookup_rel = lookup_rel let lookup_constant = lookup_constant let lookup_mind = lookup_mind let lookup_mind_specif = lookup_mind_specif @@ -248,21 +249,27 @@ let lookup_mind_specif = lookup_mind_specif (* Insertion of variables (named and de Bruijn'ed). They are now typed before being added to the environment. *) -let push_rel_or_var push (id,c) env = +let push_rel_or_var_def push (id,c) env = let (j,cst) = safe_machine env c in let env' = add_constraints cst env in - let var = assumption_of_judgment env' Evd.empty j in - push (id,var) env' + push (id,j.uj_val,j.uj_type) env' + +let push_var_def nvar env = push_rel_or_var_def push_var_def nvar env + +let push_rel_def nrel env = push_rel_or_var_def push_rel_def nrel env -let push_var nvar env = push_rel_or_var push_var nvar env +let push_rel_or_var_decl push (id,c) env = + let (j,cst) = safe_machine_type env c in + let env' = add_constraints cst env in + let var = assumption_of_type_judgment j in + push (id,var) env' -let push_rel nrel env = push_rel_or_var push_rel nrel env +let push_var_decl nvar env = push_rel_or_var_decl push_var_decl nvar env -let push_vars vars env = - List.fold_left (fun env nvar -> push_var nvar env) env vars +let push_rel_decl nrel env = push_rel_or_var_decl push_rel_decl nrel env -let push_rels vars env = - List.fold_left (fun env nvar -> push_rel nvar env) env vars +let push_rels_with_univ vars env = + List.fold_left (fun env nvar -> push_rel_decl nvar env) env vars (* Insertion of constants and parameters in environment. *) @@ -291,7 +298,7 @@ let add_constant_with_value sp body typ env = const_kind = kind_of_path sp; const_body = Some (ref (Cooked body)); const_type = ty; - const_hyps = keep_hyps (var_context env) ids; + const_hyps = keep_hyps ids (var_context env); const_constraints = Constraint.union cst cst'; const_opaque = false } in @@ -305,7 +312,7 @@ let add_lazy_constant sp f t env = const_kind = kind_of_path sp; const_body = Some (ref (Recipe f)); const_type = assumption_of_judgment env' Evd.empty jt; - const_hyps = keep_hyps (var_context env) ids; + const_hyps = keep_hyps ids (var_context env); const_constraints = cst; const_opaque = false } in @@ -327,7 +334,7 @@ let add_parameter sp t env = const_kind = kind_of_path sp; const_body = None; const_type = assumption_of_judgment env' Evd.empty jt; - const_hyps = keep_hyps (var_context env) ids; + const_hyps = keep_hyps ids (var_context env); const_constraints = cst; const_opaque = false } in @@ -344,24 +351,21 @@ let max_universe (s1,cst1) (s2,cst2) g = | Type u1, _ -> s1, cst1 | _, _ -> s2, cst2 -(* This (re)computes informations relevant to extraction and the sort of - an arity or type constructor *) +(* This (re)computes informations relevant to extraction and the sort of an + arity or type constructor; we do not to recompute universes constraints *) let rec infos_and_sort env t = match kind_of_term t with | IsProd (name,c1,c2) -> - let (varj,cst1) = safe_machine_type env c1 in + let (varj,_) = safe_machine_type env c1 in let var = assumption_of_type_judgment varj in - let env1 = Environ.push_rel (name,var) env in - let (infos,smax,cst) = infos_and_sort env1 c2 in + let env1 = Environ.push_rel_decl (name,var) env in let s1 = varj.utj_type in - let smax',cst' = max_universe (s1,cst1) (smax,cst) (universes env) in let logic = not (is_info_type env Evd.empty varj) in let small = is_small s1 in - ((logic,small) :: infos, smax', cst') + (logic,small) :: (infos_and_sort env1 c2) | IsCast (c,_) -> infos_and_sort env c - | _ -> - ([],prop (* = neutral element of max_universe *),Constraint.empty) + | _ -> [] (* [infos] is a sequence of pair [islogic,issmall] for each type in the product of a constructor or arity *) @@ -376,10 +380,12 @@ let is_unit arinfos constrsinfos = | [constrinfos] -> is_logic_constr constrinfos && is_logic_arity arinfos | _ -> false -let small_unit constrsinfos (env_par,nparams,ar) = +let small_unit constrsinfos (env_ar,nparams,ar) = let issmall = List.for_all is_small constrsinfos in - let (arinfos,_,_) = - let (_,a) = mind_extract_params nparams ar in infos_and_sort env_par a in + let arinfos = + let (params,a) = mind_extract_params nparams ar in + let env_par = push_rels params env_ar in + infos_and_sort env_par a in let isunit = is_unit arinfos constrsinfos in issmall, isunit @@ -391,14 +397,19 @@ let enforce_type_constructor arsort smax cst = | _,_ -> cst let type_one_constructor env_ar nparams arsort c = - let (infos,max,cst1) = + let infos = let (params,dc) = mind_extract_params nparams c in let env_par = push_rels params env_ar in infos_and_sort env_par dc in - let (j,cst2) = safe_machine_type env_ar c in - (*C'est idiot, cst1 et cst2 sont essentiellement des copies l'un de l'autre*) - let cst3 =enforce_type_constructor arsort max (Constraint.union cst1 cst2) in - (infos, assumption_of_type_judgment j, cst3) + (* Constructors are typed-checked here *) + let (j,cst) = safe_machine_type env_ar c in + + (* If the arity is at some level Type arsort, then the sort of the + constructor must be below arsort; here we consider constructors with the + global parameters (which add a priori more constraints on their sort) *) + let cst2 = enforce_type_constructor arsort j.utj_type cst in + + (infos, assumption_of_type_judgment j, cst2) let type_one_inductive i env_ar env_par nparams ninds (id,ar,cnames,vc) = let arsort = sort_of_arity env_ar ar in @@ -412,7 +423,7 @@ let type_one_inductive i env_ar env_par nparams ninds (id,ar,cnames,vc) = in let vc' = Array.of_list jlc in let issmall,isunit = small_unit constrsinfos (env_par,nparams,ar) in - let (_,tyar) = lookup_rel (ninds+1-i) env_ar in + let (_,tyar) = lookup_rel_type (ninds+1-i) env_ar in ((id,tyar,cnames,issmall,isunit,vc'), cst) let add_mind sp mie env = @@ -425,7 +436,8 @@ let add_mind sp mie env = let types_sign = List.map (fun (id,ar,_,_) -> (Name id,ar)) mie.mind_entry_inds in - let env_arities = push_rels types_sign env in + (* Arities with params are typed-checked here *) + let env_arities = push_rels_with_univ types_sign env in let env_params = push_rels params env_arities in let _,inds,cst = List.fold_right @@ -445,18 +457,10 @@ let add_mind sp mie env = let add_constraints = add_constraints -let pop_vars idl env = - let rec remove n sign = - if n = 0 then - sign - else - if isnull_sign sign then anomaly "pop_vars" - else - let (id,_) = hd_sign sign in - if not (List.mem id idl) then anomaly "pop_vars"; - remove (pred n) (tl_sign sign) - in - change_hyps (remove (List.length idl)) env +let rec pop_vars idl env = + match idl with + | [] -> env + | id::l -> pop_vars l (Environ.pop_var id env) let export = export let import = import diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli index 5f6b9b073..6d8c80bd0 100644 --- a/kernel/safe_typing.mli +++ b/kernel/safe_typing.mli @@ -26,8 +26,13 @@ val universes : safe_environment -> universes val context : safe_environment -> context val var_context : safe_environment -> var_context -val push_var : identifier * constr -> safe_environment -> safe_environment -val push_rel : name * constr -> safe_environment -> safe_environment +val push_var_decl : identifier * constr -> safe_environment -> safe_environment +val push_var_def : identifier * constr -> safe_environment -> safe_environment +(* +val push_rel_decl : name * constr -> safe_environment -> safe_environment +val push_rel_def : name * constr -> safe_environment -> safe_environment +*) + val add_constant : section_path -> constant_entry -> safe_environment -> safe_environment val add_lazy_constant : @@ -42,8 +47,10 @@ val add_constraints : constraints -> safe_environment -> safe_environment val pop_vars : identifier list -> safe_environment -> safe_environment -val lookup_var : identifier -> safe_environment -> name * typed_type +val lookup_var : identifier -> safe_environment -> constr option * typed_type +(* val lookup_rel : int -> safe_environment -> name * typed_type +*) val lookup_constant : section_path -> safe_environment -> constant_body val lookup_mind : section_path -> safe_environment -> mutual_inductive_body val lookup_mind_specif : inductive -> safe_environment -> inductive_instance diff --git a/kernel/sign.ml b/kernel/sign.ml index fb2c8d31d..edee43885 100644 --- a/kernel/sign.ml +++ b/kernel/sign.ml @@ -8,251 +8,190 @@ open Term (* Signatures *) -type 'a signature = identifier list * 'a list -type 'a db_signature = (name * 'a) list -type ('a,'b) sign = ENVIRON of 'a signature * 'b db_signature - -let gLOB hyps = ENVIRON (hyps,[]) - -let ids_of_sign (idl,_) = idl -let vals_of_sign (_,vals) = vals -let add_sign (id,ty) (idl,tyl) = (id::idl,ty::tyl) -let sign_it f (idl,tyl) e = List.fold_right2 f idl tyl e -let it_sign f e (idl,tyl) = List.fold_left2 f e idl tyl -let nil_sign = ([],[]) -let rev_sign (idl,tyl) = (List.rev idl, List.rev tyl) -let map_sign_typ f (idl,tyl) = (idl, List.map f tyl) -let concat_sign (idl1,tyl1) (idl2,tyl2) = (idl1@idl2, tyl1@tyl2) -let diff_sign (idl1,tyl1) (idl2,tyl2) = - (list_subtract idl1 idl2, list_subtract tyl1 tyl2) -let nth_sign (idl,tyl) n = (List.nth idl (n-1), List.nth tyl (n-1)) -let map_sign_graph f (ids,tys) = List.map2 f ids tys - -let isnull_sign = function - | ([],[]) -> true - | (_::_,_::_) -> false - | _ -> invalid_arg "isnull_sign" - -let hd_sign = function - | (id::_,ty::_) -> (id,ty) - | _ -> failwith "hd_sign" - -let tl_sign = function - | (_::idl,_::tyl) -> (idl,tyl) - | _ -> failwith "tl_sign" - -let lookup_sign id (dom,rang) = - let rec aux = function - | ([], []) -> raise Not_found - | ((id'::id'l), (ty::tyl)) -> if id' = id then (id',ty) else aux (id'l,tyl) - | _ -> anomaly "Names: lookup_sign" +let add d sign = d::sign +let map f = List.map (fun (na,c,t) -> (na,option_app f c,typed_app f t)) + +let add_decl (n,t) sign = (n,None,t)::sign +let add_def (n,c,t) sign = (n,Some c,t)::sign + +type var_declaration = identifier * constr option * typed_type +type var_context = var_declaration list + +let add_var = add +let add_var_decl = add_decl +let add_var_def = add_def +let rec lookup_id_type id = function + | (id',c,t) :: _ when id=id' -> t + | _ :: sign -> lookup_id_type id sign + | [] -> raise Not_found +let rec lookup_id_value id = function + | (id',c,t) :: _ when id=id' -> c + | _ :: sign -> lookup_id_value id sign + | [] -> raise Not_found +let rec lookup_id id = function + | (id',c,t) :: _ when id=id' -> (c,t) + | _ :: sign -> lookup_id id sign + | [] -> raise Not_found +let empty_var_context = [] +let pop_var id = function + | (id',_,_) :: sign -> assert (id = id'); sign + | [] -> assert false +let ids_of_var_context = List.map (fun (id,_,_) -> id) +let map_var_context = map +let rec mem_var_context id = function + | (id',_,_) :: _ when id=id' -> true + | _ :: sign -> mem_var_context id sign + | [] -> false +let fold_var_context = List.fold_right +let fold_var_context_both_sides = list_fold_right_and_left +let it_var_context_quantifier f = List.fold_left (fun c d -> f d c) + +type rel_declaration = name * constr option * typed_type +type rel_context = rel_declaration list + +let add_rel = add +let add_rel_decl = add_decl +let add_rel_def = add_def +let lookup_rel_type n sign = + let rec lookrec = function + | (1, (na,_,t) :: _) -> (na,t) + | (n, _ :: sign) -> lookrec (n-1,sign) + | (_, []) -> raise Not_found in - aux (dom,rang) - -let list_of_sign (ids,tys) = - try List.combine ids tys - with _ -> anomaly "Corrupted signature" - - -let make_sign = List.split -let do_sign f (ids,tys) = List.iter2 f ids tys - -let uncons_sign = function - | (id::idl,ty::tyl) -> ((id,ty),(idl,tyl)) - | _ -> anomaly "signatures are being manipulated in a non-abstract way" - -let sign_length (idl,tyl) = - let lenid = List.length idl - and lenty = List.length tyl in - if lenid = lenty then - lenid - else - invalid_arg "lookup_sign" - -let mem_sign sign id = List.mem id (ids_of_sign sign) - -let modify_sign id ty = - let rec modrec = function - | [],[] -> invalid_arg "modify_sign" - | sign -> - let (id',ty') = hd_sign sign in - if id = id' then - add_sign (id,ty) (tl_sign sign) - else - add_sign (id',ty') (modrec (tl_sign sign)) + lookrec (n,sign) +let lookup_rel_value n sign = + let rec lookrec = function + | (1, (_,c,_) :: _) -> c + | (n, _ :: sign ) -> lookrec (n-1,sign) + | (_, []) -> raise Not_found in - modrec - -let exists_sign f = - let rec exrec sign = - if isnull_sign sign then - false - else - let ((id,t),tl) = uncons_sign sign in - f id t or exrec tl + lookrec (n,sign) +let rec lookup_rel_id id sign = + let rec lookrec = function + | (n, (Anonymous,_,_)::l) -> lookrec (n+1,l) + | (n, (Name id',_,t)::l) -> if id' = id then (n,t) else lookrec (n+1,l) + | (_, []) -> raise Not_found in - exrec - -(* [sign_prefix id sign] returns the signature up to and including id, - with all later assumptions stripped off. It is an error to call it - with a signature not containing id, and that error is generated - with error. *) - -let sign_prefix id sign = - let rec prefrec sign = - if isnull_sign sign then - error "sign_prefix" - else - let ((id',t),sign') = uncons_sign sign in - if id' = id then sign else prefrec sign' + lookrec (1,sign) +let empty_rel_context = [] +let rel_context_length = List.length +let lift_rel_context n sign = + let rec liftrec k = function + | (na,c,t)::sign -> + (na,option_app (liftn n k) c,typed_app (liftn n k) t) + ::(liftrec (k-1) sign) + | [] -> [] + in + liftrec (rel_context_length sign) sign +let concat_rel_context = (@) +let ids_of_rel_context sign = + List.fold_right + (fun (na,_,_) l -> match na with Name id -> id::l | Anonymous -> l) + sign [] +let names_of_rel_context = List.map (fun (na,_,_) -> na) +let decls_of_rel_context sign = + List.fold_right + (fun (na,c,t) l -> match c with Some _ -> l | None -> (na,body_of_type t)::l) + sign [] +let map_rel_context = map + +let instantiate_sign sign args = + let rec instrec = function + | ((id,None,_) :: sign, c::args) -> (id,c) :: (instrec (sign,args)) + | ((id,Some c,_) :: _, args) -> (id,c) :: (instrec (sign,args)) + | ([],[]) -> [] + | ([],_) | (_,[]) -> + anomaly "Signature and its instance do not match" in - prefrec sign - -let add_sign_after whereid (id,t) sign = - let rec addrec sign = - if isnull_sign sign then - error "add_sign_after" - else - let ((id',t'),sign') = uncons_sign sign in - if id' = whereid then add_sign (id,t) sign - else add_sign (id',t') (addrec sign') + instrec (sign,args) + +(* [keep_hyps sign ids] keeps the part of the signature [sign] which + contains the variables of the set [ids], and recursively the variables + contained in the types of the needed variables. *) + +let rec keep_hyps needed = function + | (id,copt,t as d) ::sign when Idset.mem id needed -> + let globc = + match copt with + | None -> Idset.empty + | Some c -> global_vars_set c in + let needed' = + Idset.union (global_vars_set (body_of_type t)) + (Idset.union globc needed) in + d :: (keep_hyps needed' sign) + | _::sign -> keep_hyps needed sign + | [] -> [] + +(*************************) +(* Names environments *) +(*************************) +type names_context = name list +let add_name n nl = n::nl +let lookup_name_of_rel p names = + try List.nth names (p-1) + with Failure "nth" -> raise Not_found +let rec lookup_rel_of_name id names = + let rec lookrec n = function + | Anonymous :: l -> lookrec (n+1) l + | (Name id') :: l -> if id' = id then n else lookrec (n+1) l + | [] -> raise Not_found in - addrec sign - -let add_sign_replacing whereid (id,t) sign = - let rec addrec sign = - if isnull_sign sign then - error "add_replacing_after" - else - let ((id',t'),sign') = uncons_sign sign in - if id' = whereid then add_sign (id,t) sign' - else add_sign (id',t') (addrec sign') + lookrec 1 names +let empty_names_context = [] + +(*********************************) +(* Term destructors *) +(*********************************) + +(* Transforms a product term (x1:T1)..(xn:Tn)T into the pair + ([(xn,Tn);...;(x1,T1)],T), where T is not a product *) +let decompose_prod_assum = + let rec prodec_rec l = function + | DOP2(Prod,t,DLAM(x,c)) -> + prodec_rec (add_rel_decl (x,outcast_type t) l) c +(* | Letin,t,DLAM(x,c)) -> + prodec_rec (add_rel_def (x,c,outcast_type t) l) c*) + | DOP2(Cast,c,_) -> prodec_rec l c + | c -> l,c in - addrec sign - -(* [prepend_sign Gamma1 Gamma2] - prepends Gamma1 to the front of Gamma2, given that their namespaces - are distinct. *) - -let prepend_sign gamma1 gamma2 = - if [] = list_intersect (ids_of_sign gamma1) (ids_of_sign gamma2) then - let (ids1,vals1) = gamma1 - and (ids2,vals2) = gamma2 in - (ids1@ids2, vals1@vals2) - else - invalid_arg "prepend_sign" - -let dunbind default sign ty = function - | DLAM(na,c) -> - let id = next_ident_away (id_of_name default na) (ids_of_sign sign) in - (add_sign (id,ty) sign, subst1 (VAR id) c) - | _ -> invalid_arg "dunbind default" - -let dunbindv default sign ty = function - | DLAMV(na,c) -> - let id = next_ident_away (id_of_name default na) (ids_of_sign sign) in - (add_sign (id,ty) sign,Array.map (subst1 (VAR id)) c) - | _ -> invalid_arg "dunbindv default" - -let dbind sign c = - let (id,ty) = hd_sign sign - and sign' = tl_sign sign in - (ty,DLAM(Name id,subst_var id c)) - -let dbindv sign cl = - let (id,ty) = hd_sign sign - and sign' = tl_sign sign in - (ty,DLAMV(Name id,Array.map (subst_var id) cl)) - -(* de Bruijn environments *) - -let nil_dbsign = [] - -(* Signatures + de Bruijn environments *) - -let dbenv_it f (ENVIRON(_,dbs)) init = - List.fold_right (fun (na,t) v -> f na t v) dbs init - -let it_dbenv f init (ENVIRON(_,dbs)) = - List.fold_left (fun v (na,t) -> f v na t) init dbs - -let isnull_rel_env (ENVIRON(_,dbs)) = (dbs = []) -let uncons_rel_env (ENVIRON(sign,dbs)) = List.hd dbs,ENVIRON(sign,List.tl dbs) - -let ids_of_env (ENVIRON(sign,dbenv)) = - let filter (n,_) l = match n with (Name id) -> id::l | Anonymous -> l in - (ids_of_sign sign) @ (List.fold_right filter dbenv []) - -let get_globals (ENVIRON(g,_)) = g -let get_rels (ENVIRON(_,r)) = r - -let add_rel (n,x) (ENVIRON(g,r)) = (ENVIRON(g,(n,x)::r)) - -let add_glob (id,x) (ENVIRON((dom,rang),r)) = (ENVIRON((id::dom,x::rang),r)) - -let lookup_glob id (ENVIRON((dom,rang),_)) = - let rec aux = function - | ([], []) -> raise Not_found - | ((id'::id'l), (ty::tyl)) -> if id' = id then (id',ty) else aux (id'l,tyl) - | _ -> anomaly "Names: lookup_glob" + prodec_rec empty_rel_context + +(* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair + ([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *) +let decompose_lam_assum = + let rec lamdec_rec l = function + | DOP2(Lambda,t,DLAM(x,c)) -> + lamdec_rec (add_rel_decl (x,outcast_type t) l) c +(* | Letin,t,DLAM(x,c)) -> + lamdec_rec (add_rel_def (x,c,outcast_type t) l) c*) + | DOP2(Cast,c,_) -> lamdec_rec l c + | c -> l,c in - aux (dom,rang) - -let mem_glob id (ENVIRON((dom,_),_)) = List.mem id dom - -let lookup_rel n (ENVIRON(_,r)) = - let rec lookrec n l = match (n,l) with - | (1, ((na,x)::l)) -> (na,x) - | (n, (_::l)) -> lookrec (n-1) l - | (_, []) -> raise Not_found + lamdec_rec empty_rel_context + +(* Given a positive integer n, transforms a product term (x1:T1)..(xn:Tn)T + into the pair ([(xn,Tn);...;(x1,T1)],T) *) +let decompose_prod_n_assum n = + if n < 0 then error "decompose_prod_n: integer parameter must be positive"; + let rec prodec_rec l n c = + if n=0 then l,c + else match c with + | DOP2(Prod,t,DLAM(x,c)) -> prodec_rec (add_rel_decl (x,outcast_type t) l) (n-1) c + | DOP2(Cast,c,_) -> prodec_rec l n c + | c -> error "decompose_prod_n: not enough products" in - lookrec n r - -let rec lookup_rel_id id (ENVIRON(_,r)) = - let rec lookrec = function - | (n, ((Anonymous,x)::l)) -> lookrec (n+1,l) - | (n, ((Name id',x)::l)) -> if id' = id then (n,x) else lookrec (n+1,l) - | (_, []) -> raise Not_found + prodec_rec empty_rel_context n + +(* Given a positive integer n, transforms a lambda term [x1:T1]..[xn:Tn]T + into the pair ([(xn,Tn);...;(x1,T1)],T) *) +let decompose_lam_n_assum n = + if n < 0 then error "decompose_lam_n: integer parameter must be positive"; + let rec lamdec_rec l n c = + if n=0 then l,c + else match c with + | DOP2(Lambda,t,DLAM(x,c)) -> lamdec_rec (add_rel_decl (x,outcast_type t) l) (n-1) c + | DOP2(Cast,c,_) -> lamdec_rec l n c + | c -> error "decompose_lam_n: not enough abstractions" in - lookrec (1,r) - -let map_rel_env f (ENVIRON(g,r)) = - ENVIRON (g,List.map (fun (na,x) -> (na,f x)) r) - -let map_var_env f (ENVIRON((dom,rang),r)) = - ENVIRON (List.fold_right2 - (fun na x (doml,rangl) -> (na::doml,(f x)::rangl)) - dom rang ([],[]),r) - -let number_of_rels (ENVIRON(_,db)) = List.length db - -let change_name_env (ENVIRON(sign,db)) j id = - match list_chop (j-1) db with - | db1,((_,ty)::db2) -> ENVIRON(sign,db1@(Name id,ty)::db2) - | _ -> raise Not_found - -let unitize_env env = map_rel_env (fun _ -> ()) env - -type ('b,'a) search_result = - | GLOBNAME of identifier * 'b - | RELNAME of int * 'a - -let lookup_id id env = - try - let (x,y) = lookup_rel_id id env in RELNAME(x,y) - with - | Not_found -> let (x,y) = lookup_glob id env in GLOBNAME(x,y) - -let make_all_name_different (ENVIRON((dom,_) as sign,dbs)) = - let _,newdbs = - List.fold_right - (fun (na,t) (avoid,newdbs) -> - let id = next_name_away na avoid in - (id::avoid),((Name id,t)::newdbs)) - dbs (dom,[]) - in ENVIRON (sign,newdbs) - -type 'b assumptions = (typed_type,'b) sign -type context = (typed_type,typed_type) sign -type var_context = typed_type signature + lamdec_rec empty_rel_context n diff --git a/kernel/sign.mli b/kernel/sign.mli index 763a67471..3add55894 100644 --- a/kernel/sign.mli +++ b/kernel/sign.mli @@ -7,107 +7,75 @@ open Generic open Term (*i*) -(* Signatures of named variables. *) - -type 'a signature - -val nil_sign : 'a signature -val add_sign : (identifier * 'a) -> 'a signature -> 'a signature -val lookup_sign : identifier -> 'a signature -> (identifier * 'a) - -val rev_sign : 'a signature -> 'a signature -val map_sign_typ : ('a -> 'b) -> 'a signature -> 'b signature -val isnull_sign : 'a signature -> bool -val hd_sign : 'a signature -> identifier * 'a -val tl_sign : 'a signature -> 'a signature - -(* [sign_it f sign a] iters [f] on [sign] starting from [a] and - peeling [sign] from the oldest declaration *) - -val sign_it : (identifier -> 'a -> 'b -> 'b) -> 'a signature -> 'b -> 'b - -(* [it_sign f a sign] iters [f] on [sign] starting from [a] and - peeling [sign] from the more recent declaration *) - -val it_sign : ('b -> identifier -> 'a -> 'b) -> 'b -> 'a signature -> 'b - -val concat_sign : 'a signature -> 'a signature -> 'a signature - -val ids_of_sign : 'a signature -> identifier list -val vals_of_sign : 'a signature -> 'a list - -val nth_sign : 'a signature -> int -> (identifier * 'a) -val map_sign_graph : (identifier -> 'a -> 'b) -> 'a signature -> 'b list -val list_of_sign : 'a signature -> (identifier * 'a) list -val make_sign : (identifier * 'a) list -> 'a signature -val do_sign : (identifier -> 'a -> unit) -> 'a signature -> unit -val uncons_sign : 'a signature -> (identifier * 'a) * 'a signature -val sign_length : 'a signature -> int -val mem_sign : 'a signature -> identifier -> bool -val modify_sign : identifier -> 'a -> 'a signature -> 'a signature - -val exists_sign : (identifier -> 'a -> bool) -> 'a signature -> bool -val sign_prefix : identifier -> 'a signature -> 'a signature -val add_sign_after : - identifier -> (identifier * 'a) -> 'a signature -> 'a signature -val add_sign_replacing : - identifier -> (identifier * 'a) -> 'a signature -> 'a signature -val prepend_sign : 'a signature -> 'a signature -> 'a signature - -val dunbind : identifier -> 'a signature -> 'a -> 'b term - -> 'a signature * 'b term -val dunbindv : identifier -> 'a signature -> 'a -> 'b term - -> 'a signature * 'b term array -val dbind : 'a signature -> 'b term -> 'a * 'b term -val dbindv : 'a signature -> 'b term array -> 'a * 'b term - -(*s Signatures with named and de Bruijn variables. *) - -type 'a db_signature -type ('a,'b) sign = ENVIRON of 'a signature * 'b db_signature - -val gLOB : 'b signature -> ('b,'a) sign - -val add_rel : (name * 'a) -> ('b,'a) sign -> ('b,'a) sign -val add_glob : (identifier * 'b) -> ('b,'a) sign -> ('b,'a) sign -val lookup_glob : identifier -> ('b,'a) sign -> (identifier * 'b) -val lookup_rel : int -> ('b,'a) sign -> (name * 'a) -val mem_glob : identifier -> ('b,'a) sign -> bool - -val nil_dbsign : 'a db_signature -val get_globals : ('b,'a) sign -> 'b signature -val get_rels : ('b,'a) sign -> 'a db_signature -val dbenv_it : (name -> 'b -> 'c -> 'c) -> ('a,'b) sign -> 'c -> 'c -val it_dbenv : ('c -> name -> 'b -> 'c) -> 'c -> ('a,'b) sign -> 'c -val map_rel_env : ('a -> 'b) -> ('c,'a) sign -> ('c,'b) sign -val map_var_env : ('c -> 'b) -> ('c,'a) sign -> ('b,'a) sign -val isnull_rel_env : ('a,'b) sign -> bool -val uncons_rel_env : ('a,'b) sign -> (name * 'b) * ('a,'b) sign -val ids_of_env : ('a, 'b) sign -> identifier list -val number_of_rels : ('b,'a) sign -> int - -(*i This is for Cases i*) -(* raise [Not_found] if the integer is out of range *) -val change_name_env: ('a, 'b) sign -> int -> identifier -> ('a, 'b) sign - -val make_all_name_different : ('a, 'b) sign -> ('a, 'b) sign - -type ('b,'a) search_result = - | GLOBNAME of identifier * 'b - | RELNAME of int * 'a - -val lookup_id : identifier -> ('b,'a) sign -> ('b,'a) search_result - - -type 'b assumptions = (typed_type,'b) sign -type context = (typed_type,typed_type) sign -type var_context = typed_type signature - -val unitize_env : 'a assumptions -> unit assumptions - - - - - - - +(*s Signatures of _ordered_ named variables, intended to be accessed by name *) + +type var_context = var_declaration list + +val add_var : identifier * constr option * typed_type -> var_context -> var_context +val add_var_decl : identifier * typed_type -> var_context -> var_context +val add_var_def : identifier * constr * typed_type -> var_context ->var_context +val lookup_id : identifier -> var_context -> constr option * typed_type +val lookup_id_type : identifier -> var_context -> typed_type +val lookup_id_value : identifier -> var_context -> constr option +val pop_var : identifier -> var_context -> var_context +val empty_var_context : var_context +val ids_of_var_context : var_context -> identifier list +val map_var_context : (constr -> constr) -> var_context -> var_context +val mem_var_context : identifier -> var_context -> bool +val fold_var_context : (var_declaration -> 'a -> 'a) -> var_context -> 'a -> 'a +val fold_var_context_both_sides : + ('a -> var_declaration -> var_context -> 'a) -> var_context -> 'a -> 'a +val it_var_context_quantifier : + (var_declaration -> constr -> constr) -> constr -> var_context -> constr +val instantiate_sign : var_context -> constr list -> (identifier * constr) list +val keep_hyps : Idset.t -> var_context -> var_context + +(*s Signatures of ordered optionally named variables, intended to be + accessed by de Bruijn indices *) + +type rel_context = rel_declaration list + +val add_rel : (name * constr option * typed_type) -> rel_context -> rel_context +val add_rel_decl : (name * typed_type) -> rel_context -> rel_context +val add_rel_def : (name * constr * typed_type) -> rel_context -> rel_context +val lookup_rel_type : int -> rel_context -> name * typed_type +val lookup_rel_value : int -> rel_context -> constr option +val lookup_rel_id : identifier -> rel_context -> int * typed_type +val empty_rel_context : rel_context +val rel_context_length : rel_context -> int +val lift_rel_context : int -> rel_context -> rel_context +val concat_rel_context : newer:rel_context -> older:rel_context -> rel_context +val ids_of_rel_context : rel_context -> identifier list +val decls_of_rel_context : rel_context -> (name * constr) list +val map_rel_context : (constr -> constr) -> rel_context -> rel_context + +(*s This is used to translate names into de Bruijn indices and + vice-versa without to care about typing information *) + +type names_context +val add_name : name -> names_context -> names_context +val lookup_name_of_rel : int -> names_context -> name +val lookup_rel_of_name : identifier -> names_context -> int +val names_of_rel_context : rel_context -> names_context +val empty_names_context : names_context + +(*s Term destructors *) + +(* Transforms a product term $(x_1:T_1)..(x_n:T_n)T$ including letins + into the pair $([(x_n,T_n);...;(x_1,T_1)],T)$, where $T$ is not a + product nor a let. *) +val decompose_prod_assum : constr -> rel_context * constr + +(* Transforms a lambda term $[x_1:T_1]..[x_n:T_n]T$ including letins + into the pair $([(x_n,T_n);...;(x_1,T_1)],T)$, where $T$ is not a + lambda nor a let. *) +val decompose_lam_assum : constr -> rel_context * constr + +(* Given a positive integer n, transforms a product term + $(x_1:T_1)..(x_n:T_n)T$ + into the pair $([(xn,Tn);...;(x1,T1)],T)$. *) +val decompose_prod_n_assum : int -> constr -> rel_context * constr + +(* Given a positive integer $n$, transforms a lambda term + $[x_1:T_1]..[x_n:T_n]T$ into the pair $([(x_n,T_n);...;(x_1,T_1)],T)$ *) +val decompose_lam_n_assum : int -> constr -> rel_context * constr diff --git a/kernel/term.ml b/kernel/term.ml index 8b5b05679..6b3aa5bec 100644 --- a/kernel/term.ml +++ b/kernel/term.ml @@ -107,6 +107,9 @@ let outcast_type x = x let typed_combine f g t u = f t u (**) +type var_declaration = identifier * constr option * typed_type +type rel_declaration = name * constr option * typed_type + (****************************************************************************) (* Functions for dealing with constr terms *) (****************************************************************************) @@ -159,16 +162,53 @@ let mkCast t1 t2 = (* Constructs the product (x:t1)t2 *) let mkProd x t1 t2 = (DOP2 (Prod,t1,(DLAM (x,t2)))) - -(* non-dependent product t1 -> t2 *) -let mkArrow t1 t2 = mkProd Anonymous t1 t2 - -(* named product *) -let mkNamedProd name typ c = mkProd (Name name) typ (subst_var name c) +let mkNamedProd id typ c = mkProd (Name id) typ (subst_var id c) +let mkProd_string s t c = mkProd (Name (id_of_string s)) t c (* Constructs the abstraction [x:t1]t2 *) let mkLambda x t1 t2 = (DOP2 (Lambda,t1,(DLAM (x,t2)))) -let mkNamedLambda name typ c = mkLambda (Name name) typ (subst_var name c) +let mkNamedLambda id typ c = mkLambda (Name id) typ (subst_var id c) +let mkLambda_string s t c = mkLambda (Name (id_of_string s)) t c + +(* Constructs [x=c_1:t]c_2 *) +let mkLetIn x c1 t c2 = failwith "TODO" +let mkNamedLetIn id c1 t c2 = mkLetIn (Name id) c1 t (subst_var id c2) + +(* Constructs either [(x:t)c] or [[x=b:t]c] *) +let mkProd_or_LetIn (na,body,t) c = + match body with + | None -> mkProd na t c + | Some b -> mkLetIn na b t c + +let mkNamedProd_or_LetIn (id,body,t) c = + match body with + | None -> mkNamedProd id (body_of_type t) c + | Some b -> mkNamedLetIn id b (body_of_type t) c + +(* Constructs either [[x:t]c] or [[x=b:t]c] *) +let mkLambda_or_LetIn (na,body,t) c = + match body with + | None -> mkLambda na t c + | Some b -> mkLetIn na b t c + +let mkNamedLambda_or_LetIn (id,body,t) c = + match body with + | None -> mkNamedLambda id (body_of_type t) c + | Some b -> mkNamedLetIn id b (body_of_type t) c + +(* Constructs either [(x:t)c] or [c] where [x] is replaced by [b] *) +let mkProd_wo_LetIn (na,body,t) c = + match body with + | None -> mkProd na (body_of_type t) c + | Some b -> subst1 b c + +let mkNamedProd_wo_LetIn (id,body,t) c = + match body with + | None -> mkNamedProd id (body_of_type t) c + | Some b -> subst1 b (subst_var id c) + +(* non-dependent product t1 -> t2 *) +let mkArrow t1 t2 = mkProd Anonymous t1 t2 (* If lt = [t1; ...; tn], constructs the application (t1 ... tn) *) let mkAppL a = DOPN (AppL, a) @@ -179,7 +219,7 @@ let mkAppList a l = DOPN (AppL, Array.of_list (a::l)) let mkConst (sp,a) = DOPN (Const sp, a) (* Constructs an existential variable *) -let mkEvar n a = DOPN (Evar n, a) +let mkEvar (n,a) = DOPN (Evar n, a) (* Constructs an abstract object *) let mkAbst sp a = DOPN (Abst sp, a) @@ -215,9 +255,6 @@ let mkMutCaseA ci p c ac = with fn [ctxn] = bn. where the lenght of the jth context is ij. - - Warning: as an invariant the ti are casted during the Fix formation; - these casts are then used by destFix *) (* Here, we assume the body already constructed *) @@ -509,10 +546,6 @@ let destCase = function where the lenght of the jth context is ij. *) -let out_fixcast = function - | DOP2 (Cast, b, DOP0 (Sort s)) as c -> outcast_type c - | _ -> anomaly "destFix: malformed recursive definition" - let destGralFix a = let nbofdefs = Array.length a in let types = Array.sub a 0 (nbofdefs-1) in @@ -529,26 +562,13 @@ let destGralFix a = let destFix = function | DOPN (Fix (recindxs,i),a) -> let (types,funnames,bodies) = destGralFix a in - ((recindxs,i),((*Array.map out_fixcast *) types,funnames,bodies)) + ((recindxs,i),(types,funnames,bodies)) | _ -> invalid_arg "destFix" let destCoFix = function | DOPN ((CoFix i),a) -> let (types,funnames,bodies) = destGralFix a in - (i,((*Array.map out_fixcast *) types,funnames,bodies)) - | _ -> invalid_arg "destCoFix" - -(* Provisoire, le temps de maitriser les cast *) -let destUntypedFix = function - | DOPN (Fix (recindxs,i),a) -> - let (types,funnames,bodies) = destGralFix a in - (recindxs,i,types,funnames,bodies) - | _ -> invalid_arg "destFix" - -let destUntypedCoFix = function - | DOPN (CoFix i,a) -> - let (types,funnames,bodies) = destGralFix a in - (i,types,funnames,bodies) + (i,(types,funnames,bodies)) | _ -> invalid_arg "destCoFix" (**********************************************************************) @@ -636,9 +656,6 @@ let abs_implicit c = mkLambda Anonymous mkImplicit c let lambda_implicit a = mkLambda (Name(id_of_string"y")) mkImplicit a let lambda_implicit_lift n a = iterate lambda_implicit n (lift n a) -let mkLambda_string s t c = mkLambda (Name (id_of_string s)) t c -let mkProd_string s t c = mkProd (Name (id_of_string s)) t c - (* prod_it b [xn:Tn;..;x1:T1] = (x1:T1)..(xn:Tn)b *) let prod_it = List.fold_left (fun c (n,t) -> mkProd n t c) @@ -754,11 +771,11 @@ let rec isArity = function ([(xn,Tn);...;(x1,T1)],T), where T is not a product *) let decompose_prod = let rec prodec_rec l = function - | DOP2(Prod,t,DLAM(x,c)) -> prodec_rec ((x,t)::l) c + | DOP2(Prod,t,DLAM(x,c)) -> prodec_rec ((x,t)::l) c | DOP2(Cast,c,_) -> prodec_rec l c | c -> l,c in - prodec_rec [] + prodec_rec [] (* Transforms a lambda term [x1:T1]..[xn:Tn]T into the pair ([(xn,Tn);...;(x1,T1)],T), where T is not a lambda *) @@ -768,7 +785,7 @@ let decompose_lam = | DOP2(Cast,c,_) -> lamdec_rec l c | c -> l,c in - lamdec_rec [] + lamdec_rec [] (* Given a positive integer n, transforms a product term (x1:T1)..(xn:Tn)T into the pair ([(xn,Tn);...;(x1,T1)],T) *) @@ -1306,13 +1323,14 @@ let rec subst_meta bl c = substitution is done the list may contain only negative occurrences that will not be substituted. *) -let subst_term_occ locs c t = - let rec substcheck except k occ c t = +let subst_term_occ_gen locs occ c t = + let except = List.for_all (fun n -> n<0) locs in + let rec substcheck k occ c t = if except or List.exists (function u -> u>=occ) locs then - substrec except k occ c t + substrec k occ c t else (occ,t) - and substrec except k occ c t = + and substrec k occ c t = if eq_constr t c then if except then if List.mem (-occ) locs then (occ+1,t) else (occ+1,Rel(k)) @@ -1327,44 +1345,60 @@ let subst_term_occ locs c t = let (occ',cl') = Array.fold_left (fun (nocc',lfd) f -> - let (nocc'',f') = substcheck except k nocc' c f in + let (nocc'',f') = substcheck k nocc' c f in (nocc'',f'::lfd)) (occ,[]) cl in (occ',DOPN(i,Array.of_list (List.rev cl'))) | DOP2(i,t1,t2) -> - let (nocc1,t1')=substrec except k occ c t1 in - let (nocc2,t2')=substcheck except k nocc1 c t2 in + let (nocc1,t1')=substrec k occ c t1 in + let (nocc2,t2')=substcheck k nocc1 c t2 in nocc2,DOP2(i,t1',t2') | DOP1(i,t) -> - let (nocc,t')= substrec except k occ c t in + let (nocc,t')= substrec k occ c t in nocc,DOP1(i,t') | DLAM(n,t) -> - let (occ',t') = substcheck except (k+1) occ (lift 1 c) t in + let (occ',t') = substcheck (k+1) occ (lift 1 c) t in (occ',DLAM(n,t')) | DLAMV(n,cl) -> let (occ',cl') = Array.fold_left (fun (nocc',lfd) f -> let (nocc'',f') = - substcheck except (k+1) nocc' (lift 1 c) f + substcheck (k+1) nocc' (lift 1 c) f in (nocc'',f'::lfd)) (occ,[]) cl in (occ',DLAMV(n,Array.of_list (List.rev cl') )) | _ -> occ,t - in - if locs = [] then + in + substcheck 1 occ c t + +let subst_term_occ locs c t = + if locs = [] then subst_term c t else if List.mem 0 locs then t else - let except = List.for_all (fun n -> n<0) locs in - let (nbocc,t') = substcheck except 1 1 c t in + let (nbocc,t') = subst_term_occ_gen locs 1 c t in if List.exists (fun o -> o >= nbocc or o <= -nbocc) locs then - failwith "subst_term_occ: too few occurences"; + errorlabstrm "subst_term_occ" [< 'sTR "Too few occurences" >]; t' +let subst_term_occ_decl locs c (id,bodyopt,typ as d) = + match bodyopt with + | None -> (id,None,subst_term_occ locs c typ) + | Some body -> + if locs = [] then + (id,Some (subst_term c body),typed_app (subst_term c) typ) + else if List.mem 0 locs then + d + else + let (nbocc,body') = subst_term_occ_gen locs 1 c body in + let (nbocc',t') = typed_app (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" >]; + (id,Some body',t') (***************************) (* occurs check functions *) diff --git a/kernel/term.mli b/kernel/term.mli index 72ba6f511..9019fa091 100644 --- a/kernel/term.mli +++ b/kernel/term.mli @@ -59,10 +59,9 @@ type constr = sorts oper term type flat_arity = (name * constr) list * sorts -type 'a judge = { body : constr; typ : 'a } +(*type 'a judge = { body : constr; typ : 'a }*) type typed_type -type typed_term = typed_type judge val make_typed : constr -> sorts -> typed_type val make_typed_lazy : constr -> (constr -> sorts) -> typed_type @@ -78,6 +77,9 @@ val incast_type : typed_type -> constr val outcast_type : constr -> typed_type +type var_declaration = identifier * constr option * typed_type +type rel_declaration = name * constr option * typed_type + (**********************************************************************) type binder_kind = BProd | BLambda @@ -96,7 +98,10 @@ type 'ctxt reference = manipulation of terms. Some of these functions may be overlapped with previous ones. *) -(* Concrete type for making pattern-matching. *) +(* Concrete type for making pattern-matching. *) + +(* [constr array] is an instance matching definitional var_context in + the same order (i.e. last argument first) *) type existential = int * constr array type constant = section_path * constr array type constructor = constructor_path * constr array @@ -169,18 +174,30 @@ val implicit_sort : sorts type $t_2$ (that means t2 is declared as the type of t1). *) val mkCast : constr -> constr -> constr -(* Constructs the product $(x:t_1)t_2$. $x$ may be anonymous. *) +(* Constructs the product $(x:t_1)t_2$ *) val mkProd : name -> constr -> constr -> constr - -(* [mkProd_string s t c] constructs the product $(s:t)c$ *) +val mkNamedProd : identifier -> constr -> constr -> constr val mkProd_string : string -> constr -> constr -> constr +(* Constructs the product $(x:t_1)t_2$ *) +val mkLetIn : name -> constr -> constr -> constr -> constr +val mkNamedLetIn : identifier -> constr -> constr -> constr -> constr + +(* Constructs either [(x:t)c] or [[x=b:t]c] *) +val mkProd_or_LetIn : rel_declaration -> constr -> constr +val mkNamedProd_or_LetIn : var_declaration -> constr -> constr + +(* Constructs either [[x:t]c] or [[x=b:t]c] *) +val mkLambda_or_LetIn : rel_declaration -> constr -> constr +val mkNamedLambda_or_LetIn : var_declaration -> constr -> constr + +(* Constructs either [(x:t)c] or [c] where [x] is replaced by [b] *) +val mkProd_wo_LetIn : rel_declaration -> constr -> constr +val mkNamedProd_wo_LetIn : var_declaration -> constr -> constr + (* non-dependant product $t_1 \rightarrow t_2$ *) val mkArrow : constr -> constr -> constr -(* named product *) -val mkNamedProd : identifier -> constr -> constr -> constr - (* Constructs the abstraction $[x:t_1]t_2$ *) val mkLambda : name -> constr -> constr -> constr val mkNamedLambda : identifier -> constr -> constr -> constr @@ -198,7 +215,7 @@ val mkAppList : constr -> constr list -> constr val mkConst : constant -> constr (* Constructs an existential variable *) -val mkEvar : int -> constr array -> constr +val mkEvar : existential -> constr (* Constructs an abstract object *) val mkAbst : section_path -> constr array -> constr @@ -360,13 +377,6 @@ val destFix : constr -> fixpoint val destCoFix : constr -> cofixpoint -(* Provisoire, le temps de maitriser les cast *) -val destUntypedFix : - constr -> int array * int * constr array * Names.name list * constr array -val destUntypedCoFix : - constr -> int * constr array * Names.name list * constr array - - (*s Other term constructors. *) val abs_implicit : constr -> constr @@ -410,7 +420,8 @@ val prod_applist : constr -> constr list -> constr (*s Other term destructors. *) (* Transforms a product term $(x_1:T_1)..(x_n:T_n)T$ into the pair - $([(x_n,T_n);...;(x_1,T_1)],T)$, where $T$ is not a product. *) + $([(x_n,T_n);...;(x_1,T_1)],T)$, where $T$ is not a product. + It includes also local definitions *) val decompose_prod : constr -> (name*constr) list * constr (* Transforms a lambda term $[x_1:T_1]..[x_n:T_n]T$ into the pair @@ -420,11 +431,11 @@ val decompose_lam : constr -> (name*constr) list * constr (* Given a positive integer n, transforms a product term $(x_1:T_1)..(x_n:T_n)T$ into the pair $([(xn,Tn);...;(x1,T1)],T)$. *) -val decompose_prod_n : int -> constr -> (name*constr) list * constr +val decompose_prod_n : int -> constr -> (name * constr) list * constr (* Given a positive integer $n$, transforms a lambda term $[x_1:T_1]..[x_n:T_n]T$ into the pair $([(x_n,T_n);...;(x_1,T_1)],T)$ *) -val decompose_lam_n : int -> constr -> (name*constr) list * constr +val decompose_lam_n : int -> constr -> (name * constr) list * constr (* [nb_lam] $[x_1:T_1]...[x_n:T_n]c$ where $c$ is not an abstraction gives $n$ (casts are ignored) *) @@ -565,11 +576,16 @@ val rename_bound_var : identifier list -> constr -> constr val eta_reduce_head : constr -> constr val eq_constr : constr -> constr -> bool val eta_eq_constr : constr -> constr -> bool -val subst_term : constr -> constr -> constr -val subst_term_eta_eq : constr -> constr -> constr + +(*s The following functions substitutes [what] by [Rel 1] in [where] *) +val subst_term : what:constr -> where:constr -> constr +val subst_term_eta_eq : what:constr -> where:constr -> constr +val subst_term_occ : occs:int list -> what:constr -> where:constr -> constr +val subst_term_occ_decl : occs:int list -> what:constr -> + where:var_declaration -> var_declaration + val replace_consts : (section_path * (identifier list * constr) option) list -> constr -> constr -val subst_term_occ : int list -> constr -> constr -> constr (* [subst_meta bl c] substitutes the metavar $i$ by $c_i$ in [c] for each binding $(i,c_i)$ in [bl], diff --git a/kernel/type_errors.ml b/kernel/type_errors.ml index d2891180a..14409c235 100644 --- a/kernel/type_errors.ml +++ b/kernel/type_errors.ml @@ -40,61 +40,59 @@ type type_error = | WrongPredicateArity of constr * int * int | NeedsInversion of constr * constr -exception TypeError of path_kind * context * type_error +exception TypeError of path_kind * env * type_error -let context_ise sigma env = - map_var_env (typed_app (Reduction.nf_ise1 sigma)) - (map_rel_env (typed_app (Reduction.nf_ise1 sigma)) - (context env)) +let env_ise sigma env = + map_context (Reduction.nf_ise1 sigma) env let error_unbound_rel k env sigma n = - raise (TypeError (k, context_ise sigma env, UnboundRel n)) + raise (TypeError (k, env_ise sigma env, UnboundRel n)) let error_not_type k env c = - raise (TypeError (k, context env, NotAType c)) + raise (TypeError (k, env, NotAType c)) let error_assumption k env c = - raise (TypeError (k, context env, BadAssumption c)) + raise (TypeError (k, env, BadAssumption c)) let error_reference_variables k env id = - raise (TypeError (k, context env, ReferenceVariables id)) + raise (TypeError (k, env, ReferenceVariables id)) let error_elim_arity k env ind aritylst c p pt okinds = - raise (TypeError (k, context env, ElimArity (ind,aritylst,c,p,pt,okinds))) + raise (TypeError (k, env, ElimArity (ind,aritylst,c,p,pt,okinds))) let error_case_not_inductive k env c ct = - raise (TypeError (k, context env, CaseNotInductive (c,ct))) + raise (TypeError (k, env, CaseNotInductive (c,ct))) let error_number_branches k env c ct expn = - raise (TypeError (k, context env, NumberBranches (c,ct,expn))) + raise (TypeError (k, env, NumberBranches (c,ct,expn))) let error_ill_formed_branch k env c i actty expty = - raise (TypeError (k, context env, IllFormedBranch (c,i,actty,expty))) + raise (TypeError (k, env, IllFormedBranch (c,i,actty,expty))) let error_generalization k env sigma nvar c = - raise (TypeError (k, context_ise sigma env, Generalization (nvar,c))) + raise (TypeError (k, env_ise sigma env, Generalization (nvar,c))) let error_actual_type k env c actty expty = - raise (TypeError (k, context env, ActualType (c,actty,expty))) + raise (TypeError (k, env, ActualType (c,actty,expty))) let error_cant_apply_not_functional k env rator randl = - raise (TypeError (k, context env, CantApplyNonFunctional (rator,randl))) + raise (TypeError (k, env, CantApplyNonFunctional (rator,randl))) let error_cant_apply_bad_type k env sigma t rator randl = - raise(TypeError (k, context_ise sigma env, CantApplyBadType (t,rator,randl))) + raise(TypeError (k, env_ise sigma env, CantApplyBadType (t,rator,randl))) let error_ill_formed_rec_body k env str lna i vdefs = - raise (TypeError (k, context env, IllFormedRecBody (str,lna,i,vdefs))) + raise (TypeError (k, env, IllFormedRecBody (str,lna,i,vdefs))) let error_ill_typed_rec_body k env i lna vdefj vargs = - raise (TypeError (k, context env, IllTypedRecBody (i,lna,vdefj,vargs))) + raise (TypeError (k, env, IllTypedRecBody (i,lna,vdefj,vargs))) let error_not_inductive k env c = - raise (TypeError (k, context env, NotInductive c)) + raise (TypeError (k, env, NotInductive c)) let error_ml_case k env mes c ct br brt = - raise (TypeError (k, context env, MLCase (mes,c,ct,br,brt))) + raise (TypeError (k, env, MLCase (mes,c,ct,br,brt))) let error_not_product env c = - raise (TypeError (CCI, context env, NotProduct c)) + raise (TypeError (CCI, env, NotProduct c)) diff --git a/kernel/type_errors.mli b/kernel/type_errors.mli index 38d6d8b6e..284f244f1 100644 --- a/kernel/type_errors.mli +++ b/kernel/type_errors.mli @@ -42,7 +42,7 @@ type type_error = | WrongPredicateArity of constr * int * int | NeedsInversion of constr * constr -exception TypeError of path_kind * context * type_error +exception TypeError of path_kind * env * type_error val error_unbound_rel : path_kind -> env -> 'a Evd.evar_map -> int -> 'b diff --git a/kernel/typeops.ml b/kernel/typeops.ml index a22e5715b..4077d852f 100644 --- a/kernel/typeops.ml +++ b/kernel/typeops.ml @@ -37,12 +37,11 @@ let type_judgment env sigma j = let assumption_of_type_judgment j = make_typed j.utj_val j.utj_type - (* Type of a de Bruijn index. *) let relative env sigma n = try - let (_,typ) = lookup_rel n env in + let (_,typ) = lookup_rel_type n env in { uj_val = Rel n; uj_type = typed_app (lift n) typ } with Not_found -> @@ -51,13 +50,13 @@ let relative env sigma n = (* Management of context of variables. *) (* Checks if a context of variable is included in another one. *) - +(* let rec hyps_inclusion env sigma sign1 sign2 = - if isnull_sign sign1 then true + if sign1 = empty_var_context then true else let (id1,ty1) = hd_sign sign1 in let rec search sign2 = - if isnull_sign sign2 then false + if sign2 = empty_var_context then false else let (id2,ty2) = hd_sign sign2 in if id1 = id2 then @@ -68,15 +67,16 @@ let rec hyps_inclusion env sigma sign1 sign2 = search (tl_sign sign2) in search sign2 +*) (* Checks if the given context of variables [hyps] is included in the current context of [env]. *) - +(* let check_hyps id env sigma hyps = let hyps' = var_context env in if not (hyps_inclusion env sigma hyps hyps') then error_reference_variables CCI env id - +*) (* Instantiation of terms on real arguments. *) let type_of_constant = Instantiate.constant_type @@ -263,8 +263,11 @@ let sort_of_product_without_univ domsort rangsort = | Prop _ -> rangsort | Type u1 -> Type dummy_univ) -let typed_product_without_universes name var c = - typed_combine (mkProd name) sort_of_product_without_univ var c +let generalize_without_universes (_,_,var as d) c = + typed_combine + (fun _ c -> mkNamedProd_or_LetIn d c) + sort_of_product_without_univ + var c let typed_product env name var c = (* Gros hack ! *) @@ -443,12 +446,11 @@ let is_subterm_specif env sigma lcx mind_recvec = stl.(0) | (DOPN(Fix(_),la) as mc,l) -> - let (recindxs,i,typarray,funnames,bodies) = destUntypedFix mc in + let ((recindxs,i),(typarray,funnames,bodies)) = destFix mc in let nbfix = List.length funnames in let decrArg = recindxs.(i) in let theBody = bodies.(i) in - let (gamma,strippedBody) = decompose_lam_n (decrArg+1) theBody in - let absTypes = List.map snd gamma in + let (_,strippedBody) = decompose_lam_n (decrArg+1) theBody in let nbOfAbst = nbfix+decrArg+1 in let newlst = if List.length l < (decrArg+1) then @@ -563,7 +565,7 @@ let rec check_subterm_rec_meta env sigma vectn k def = | (DOPN(Fix(_),la) as mc,l) -> (List.for_all (check_rec_call n lst) l) && - let (recindxs,i,typarray,funnames,bodies) = destUntypedFix mc in + let ((recindxs,i),(typarray,funnames,bodies)) = destFix mc in let nbfix = List.length funnames in let decrArg = recindxs.(i) in if (List.length l < (decrArg+1)) then @@ -831,18 +833,4 @@ let control_only_guard env sigma = in control_rec -(* [keep_hyps sign ids] keeps the part of the signature [sign] which - contains the variables of the set [ids], and recursively the variables - contained in the types of the needed variables. *) - -let keep_hyps sign needed = - rev_sign - (fst (it_sign - (fun ((hyps,globs) as sofar) id a -> - if Idset.mem id globs then - (add_sign (id,a) hyps, - Idset.union (global_vars_set (body_of_type a)) globs) - else - sofar) - (nil_sign,needed) sign)) diff --git a/kernel/typeops.mli b/kernel/typeops.mli index 6d1faa0ff..a07206fb0 100644 --- a/kernel/typeops.mli +++ b/kernel/typeops.mli @@ -51,8 +51,7 @@ val judge_of_prop_contents : contents -> unsafe_judgment val judge_of_type : universe -> unsafe_judgment * constraints -val typed_product_without_universes : - name -> typed_type -> typed_type -> typed_type +val generalize_without_universes : var_declaration -> typed_type -> typed_type val abs_rel : env -> 'a evar_map -> name -> typed_type -> unsafe_judgment @@ -85,6 +84,7 @@ open Inductive val find_case_dep_nparams : env -> 'a evar_map -> constr * constr -> inductive_family -> constr -> bool +(* val hyps_inclusion : env -> 'a evar_map -> var_context -> var_context -> bool +*) -val keep_hyps : var_context -> Idset.t -> var_context diff --git a/library/declare.ml b/library/declare.ml index 0bce97fe4..404ecda31 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -38,11 +38,15 @@ let make_strength_2 () = make_strength path -(* Variables. *) +(* Section variables. *) + +type section_variable_entry = + | SectionLocalDef of constr + | SectionLocalDecl of constr type sticky = bool -type variable_declaration = constr * strength * sticky +type variable_declaration = section_variable_entry * strength * sticky let vartab = ref (Spmap.empty : (identifier * variable_declaration) Spmap.t) @@ -51,8 +55,11 @@ let _ = Summary.declare_summary "VARIABLE" Summary.unfreeze_function = (fun ft -> vartab := ft); Summary.init_function = (fun () -> vartab := Spmap.empty) } -let cache_variable (sp,(id,(ty,_,_) as vd,imps)) = - Global.push_var (id,ty); +let cache_variable (sp,((id,(d,_,_) as vd),imps)) = + begin match d with (* Fails if not well-typed *) + | SectionLocalDecl ty -> Global.push_var_decl (id,ty) + | SectionLocalDef c -> Global.push_var_def (id,c) + end; Nametab.push id sp; if imps then declare_var_implicits id; vartab := Spmap.add sp vd !vartab @@ -63,7 +70,7 @@ let open_variable _ = () let specification_variable x = x -let (in_variable, out_variable) = +let (in_variable, _ (* out_variable *) ) = let od = { cache_function = cache_variable; load_function = load_variable; @@ -197,8 +204,8 @@ let is_variable id = let out_variable sp = let (id,(_,str,sticky)) = Spmap.find sp !vartab in - let (_,ty) = Global.lookup_var id in - (id,ty,str,sticky) + let (c,ty) = Global.lookup_var id in + ((id,c,ty),str,sticky) let variable_strength id = let sp = Nametab.sp_of_id CCI id in @@ -229,22 +236,68 @@ let mind_oper_of_id sp id mib = mip.mind_consnames) mib.mind_packets -let construct_operator env sp id = +let global_sp_operator env sp id = try let cb = Environ.lookup_constant sp env in Const sp, cb.const_hyps with Not_found -> let mib = Environ.lookup_mind sp env in mind_oper_of_id sp id mib, mib.mind_hyps -let global_operator sp id = construct_operator (Global.env()) sp id +let global_operator kind id = + let sp = Nametab.sp_of_id kind id in + global_sp_operator (Global.env()) sp id +(* let construct_sp_reference env sp id = - let (oper,hyps) = construct_operator env sp id in + let (oper,hyps) = global_sp_operator env sp id in let hyps' = Global.var_context () in if not (hyps_inclusion env Evd.empty hyps hyps') then error_reference_variables CCI env id; let ids = ids_of_sign hyps in DOPN(oper, Array.of_list (List.map (fun id -> VAR id) ids)) +*) + +let occur_decl env (id,c,t) hyps = + try + let (c',t') = Sign.lookup_id id hyps in + let matching_bodies = match c,c' with + | None, _ -> true + | Some c, None -> false + | Some c, Some c' -> is_conv env Evd.empty c c' in + let matching_types = + is_conv env Evd.empty (body_of_type t) (body_of_type t') in + matching_types & matching_bodies + with Not_found -> false + +(* +let rec find_common_hyps_then_abstract c env hyps' = function + | (id,_,_ as d) :: hyps when occur_decl env d hyps' -> + find_common_hyps_then_abstract c (Environ.push_var d env) hyps' hyps + | hyps -> + Environ.it_mkNamedLambda_or_LetIn c hyps + +let find_common_hyps_then_abstract c env hyps' hyps = + find_common_hyps_then_abstract c env hyps' (List.rev hyps) +*) + +let find_common_hyps_then_abstract c env hyps' hyps = + snd (fold_var_context_both_sides + (fun + (env,c) (id,_,_ as d) hyps -> + if occur_decl env d hyps' then + (Environ.push_var d env,c) + else + (env, Environ.it_mkNamedLambda_or_LetIn c hyps)) + hyps + (env,c)) + +let construct_sp_reference env sp id = + let (oper,hyps) = global_sp_operator env sp id in + let hyps0 = Global.var_context () in + let env0 = Environ.reset_context env in + let args = List.map mkVar (ids_of_var_context hyps) in + let body = DOPN(oper,Array.of_list args) in + find_common_hyps_then_abstract body env0 hyps0 hyps let construct_reference env kind id = try @@ -271,11 +324,11 @@ let global_reference_imps kind id = | VAR id -> c, implicits_of_var id | _ -> assert false - +(* let global env id = try let _ = lookup_glob id (Environ.context env) in VAR id with Not_found -> global_reference CCI id - +*) let is_global id = try let osp = Nametab.sp_of_id CCI id in @@ -299,10 +352,20 @@ let path_of_inductive_path (sp,tyi) = (* Eliminations. *) +let eliminations = [ (prop,"_ind") ; (spec,"_rec") ; (types,"_rect") ] + +let elimination_suffix = function + | Type _ -> "_rect" + | Prop Null -> "_ind" + | Prop Pos -> "_rec" + let declare_eliminations sp i = let oper = MutInd (sp,i) in let mib = Global.lookup_mind sp in - let ids = ids_of_sign mib.mind_hyps in + let ids = ids_of_var_context mib.mind_hyps in + if not (list_subset ids (ids_of_var_context (Global.var_context ()))) then + error ("Declarations of elimination scheme outside the section "^ + "of the inductive definition is not implemented"); let ctxt = Array.of_list (List.map (fun id -> VAR id) ids) in let mispec = Global.lookup_mind_specif ((sp,i),ctxt) in let mindstr = string_of_id (mis_typename mispec) in @@ -318,11 +381,15 @@ let declare_eliminations sp i = let npars = mis_nparams mispec in let make_elim s = instanciate_indrec_scheme s npars elim_scheme in let kelim = mis_kelim mispec in - if List.mem prop kelim then - declare (mindstr^"_ind") (make_elim prop); - if List.mem spec kelim then - declare (mindstr^"_rec") (make_elim spec); - if List.mem types kelim then - declare (mindstr^"_rect") (make_elim (Type (Univ.new_univ sp))) + List.iter + (fun (sort,suff) -> + if List.mem sort kelim then declare (mindstr^suff) (make_elim sort)) + eliminations + +(* Look up function for the default elimination constant *) + +let lookup_eliminator env path s = + let s = (string_of_id (basename path))^(elimination_suffix s) in + construct_reference env (kind_of_path path) (id_of_string s) diff --git a/library/declare.mli b/library/declare.mli index 9bb9c1607..f078c9166 100644 --- a/library/declare.mli +++ b/library/declare.mli @@ -20,9 +20,13 @@ type strength = | DischargeAt of section_path | NeverDischarge +type section_variable_entry = + | SectionLocalDef of constr + | SectionLocalDecl of constr + type sticky = bool -type variable_declaration = constr * strength * sticky +type variable_declaration = section_variable_entry * strength * sticky val declare_variable : identifier -> variable_declaration -> unit type constant_declaration = constant_entry * strength @@ -46,16 +50,15 @@ val constant_strength : section_path -> strength val constant_or_parameter_strength : section_path -> strength val is_variable : identifier -> bool -val out_variable : - section_path -> identifier * typed_type * strength * sticky +val out_variable : section_path -> var_declaration * strength * sticky val variable_strength : identifier -> strength -(*s [global_operator sp id] returns the operator (constant, inductive or - construtor) corresponding to [(sp,id)] in global environment, together +(*s [global_operator k id] returns the operator (constant, inductive or + construtor) corresponding to [id] in global environment, together with its definition environment. *) -val global_operator : section_path -> identifier -> sorts oper * var_context +val global_operator : path_kind -> identifier -> sorts oper * var_context (*s [global_reference k id] returns the object corresponding to the name [id] in the global environment. It may be a constant, @@ -76,3 +79,7 @@ val is_global : identifier -> bool val path_of_inductive_path : inductive_path -> section_path val path_of_constructor_path : constructor_path -> section_path + +(* Look up function for the default elimination constant *) +val elimination_suffix : sorts -> string +val lookup_eliminator : Environ.env -> section_path -> sorts -> constr diff --git a/library/global.ml b/library/global.ml index cf3d88ca9..a37863a78 100644 --- a/library/global.ml +++ b/library/global.ml @@ -30,8 +30,12 @@ let universes () = universes !global_env let context () = context !global_env let var_context () = var_context !global_env -let push_var idc = global_env := push_var idc !global_env -let push_rel nac = global_env := push_rel nac !global_env +let push_var_def idc = global_env := push_var_def idc !global_env +let push_var_decl idc = global_env := push_var_decl idc !global_env +(* +let push_rel_def nac = global_env := push_rel nac !global_env +let push_rel_decl nac = global_env := push_rel nac !global_env +*) let add_constant sp ce = global_env := add_constant sp ce !global_env let add_parameter sp c = global_env := add_parameter sp c !global_env let add_mind sp mie = global_env := add_mind sp mie !global_env @@ -40,7 +44,9 @@ let add_constraints c = global_env := add_constraints c !global_env let pop_vars ids = global_env := pop_vars ids !global_env let lookup_var id = lookup_var id !global_env +(* let lookup_rel n = lookup_rel n !global_env +*) let lookup_constant sp = lookup_constant sp !global_env let lookup_mind sp = lookup_mind sp !global_env let lookup_mind_specif c = lookup_mind_specif c !global_env diff --git a/library/global.mli b/library/global.mli index 19d516625..c2e8ac386 100644 --- a/library/global.mli +++ b/library/global.mli @@ -23,8 +23,13 @@ val universes : unit -> universes val context : unit -> context val var_context : unit -> var_context -val push_var : identifier * constr -> unit -val push_rel : name * constr -> unit +val push_var_decl : identifier * constr -> unit +val push_var_def : identifier * constr -> unit +(* +val push_var : identifier * constr option * constr -> unit +val push_rel_decl : name * constr -> unit +val push_rel_def : name * constr -> unit +*) val add_constant : section_path -> constant_entry -> unit val add_parameter : section_path -> constr -> unit val add_mind : section_path -> mutual_inductive_entry -> unit @@ -32,8 +37,8 @@ val add_constraints : constraints -> unit val pop_vars : identifier list -> unit -val lookup_var : identifier -> name * typed_type -val lookup_rel : int -> name * typed_type +val lookup_var : identifier -> constr option * typed_type +(*val lookup_rel : int -> name * typed_type*) val lookup_constant : section_path -> constant_body val lookup_mind : section_path -> mutual_inductive_body val lookup_mind_specif : inductive -> inductive_instance diff --git a/library/indrec.ml b/library/indrec.ml index d24ce2aa6..42d067d49 100644 --- a/library/indrec.ml +++ b/library/indrec.ml @@ -70,7 +70,7 @@ let mis_make_case_com depopt env sigma mispec kind = in let indf = make_ind_family (mispec,rel_list 0 nparams) in let typP = make_arity env' dep indf kind in - it_lambda_name env (mkLambda_string "P" typP (add_branch 0)) lnamespar + it_mkLambda_or_LetIn_name env (mkLambda_string "P" typP (add_branch 0)) lnamespar (* check if the type depends recursively on one of the inductive scheme *) @@ -154,32 +154,34 @@ let make_rec_branch_arg env sigma (nparams,fvect,decF) f cstr recargs = prec 0 in (* ici, cstrprods est la liste des produits du constructeur instantié *) - let rec process_constr i cstrprods f recargs = - match cstrprods with - | (n,t)::cprest -> - let (optionpos,rest) = - match recargs with - | [] -> (* Impossible?! *) None,[] - | (Param(i)::rest) -> None,rest - | (Norec::rest) -> None,rest - | (Imbr _::rest) -> None,rest - | (Mrec i::rest) -> fvect.(i),rest - in - (match optionpos with - | None -> - lambda_name env - (n,t,process_constr (i+1) cprest - (applist(whd_beta_stack (lift 1 f) [(Rel 1)])) rest) - | Some(_,f_0) -> - let nF = lift (i+1+decF) f_0 in - let arg = process_pos nF (lift 1 t) in - lambda_name env - (n,t,process_constr (i+1) cprest - (applist(whd_beta_stack (lift 1 f) [(Rel 1); arg])) - rest)) - | [] -> f - in - process_constr 0 (List.rev cstr.cs_args) f recargs + let rec process_constr i f = function + | (n,None,t)::cprest, recarg::rest -> + let optionpos = + match recarg with + | Param i -> None + | Norec -> None + | Imbr _ -> None + | Mrec i -> fvect.(i) + in + (match optionpos with + | None -> + lambda_name env + (n,incast_type t,process_constr (i+1) + (applist(whd_beta_stack (lift 1 f) [(Rel 1)])) + (cprest,rest)) + | Some(_,f_0) -> + let nF = lift (i+1+decF) f_0 in + let arg = process_pos nF (lift 1 (body_of_type t)) in + lambda_name env + (n,incast_type t,process_constr (i+1) + (applist(whd_beta_stack (lift 1 f) [(Rel 1); arg])) + (cprest,rest))) + | (n,Some c,t)::cprest, rest -> failwith "TODO" + | [],[] -> f + | _,[] | [],_ -> anomaly "process_constr" + + in + process_constr 0 f (List.rev cstr.cs_args, recargs) (* Main function *) let mis_make_indrec env sigma listdepkind mispec = @@ -290,7 +292,7 @@ let mis_make_indrec env sigma listdepkind mispec = if mis_is_recursive_subset (List.map (fun (mispec,_,_) -> mis_index mispec) listdepkind) mispeci then - it_lambda_name env (put_arity 0 listdepkind) lnamespar + it_mkLambda_or_LetIn_name env (put_arity 0 listdepkind) lnamespar else mis_make_case_com (Some dep) env sigma mispeci kind in diff --git a/parsing/astterm.ml b/parsing/astterm.ml index 841a571e9..002925e16 100644 --- a/parsing/astterm.ml +++ b/parsing/astterm.ml @@ -167,24 +167,32 @@ let ref_from_constr = function | VAR id -> RVar id (* utilisé pour coe_value (tmp) *) | _ -> anomaly "Not a reference" -let dbize_ref k sigma env loc s = +(* [vars1] is a set of name to avoid (used for the tactic language); + [vars2] is the set of global variables, env is the set of variables + abstracted until this point *) +let dbize_ref k sigma env loc s (vars1,vars2)= let id = ident_of_nvar loc s in + if Idset.mem id env then + RRef (loc,RVar id),[] + else + if List.mem s vars1 then + RRef(loc,RVar id),[] + else + try + let _ = lookup_id id vars2 in + RRef (loc,RVar id), (try implicits_of_var id with _ -> []) + with Not_found -> try - match lookup_id id env with - | RELNAME(n,_) -> RRef (loc,RVar id),[] - | _ -> RRef(loc,RVar id), (try implicits_of_var id with _ -> []) + let c,il = match k with + | CCI -> Declare.global_reference_imps CCI id + | FW -> Declare.global_reference_imps FW id + | OBJ -> anomaly "search_ref_cci_fw" in + RRef (loc,ref_from_constr c), il with Not_found -> - try - let c,il = match k with - | CCI -> Declare.global_reference_imps CCI id - | FW -> Declare.global_reference_imps FW id - | OBJ -> anomaly "search_ref_cci_fw" in - RRef (loc,ref_from_constr c), il - with Not_found -> - try - (Syntax_def.search_syntactic_definition id, []) - with Not_found -> - error_var_not_found_loc loc CCI id + try + (Syntax_def.search_syntactic_definition id, []) + with Not_found -> + error_var_not_found_loc loc CCI id let mkLambdaC (x,a,b) = ope("LAMBDA",[a;slam(Some (string_of_id x),b)]) let mkLambdaCit = List.fold_right (fun (x,a) b -> mkLambdaC(x,a,b)) @@ -268,22 +276,18 @@ let check_capture s ty = function let dbize k sigma env allow_soapp lvar = let rec dbrec env = function - | Nvar(loc,s) -> - if List.mem s lvar then - RRef(loc,RVar (id_of_string s)) - else - fst (dbize_ref k sigma env loc s) + | Nvar(loc,s) -> fst (dbize_ref k sigma env loc s lvar) (* | Slam(_,ona,Node(_,"V$",l)) -> let na = (match ona with Some s -> Name (id_of_string s) | _ -> Anonymous) - in DLAMV(na,Array.of_list (List.map (dbrec (add_rel (na,()) env)) l)) + in DLAMV(na,Array.of_list (List.map (dbrec (Idset.add na env)) l)) | Slam(_,ona,t) -> let na = (match ona with Some s -> Name (id_of_string s) | _ -> Anonymous) - in DLAM(na, dbrec (add_rel (na,()) env) t) + in DLAM(na, dbrec (Idset.add na env) t) *) | Node(loc,"FIX", (Nvar (locid,iddef))::ldecl) -> let (lf,ln,lA,lt) = dbize_fix ldecl in @@ -293,7 +297,7 @@ let dbize k sigma env allow_soapp lvar = with Not_found -> error_fixname_unbound "dbize (FIX)" false locid iddef in let ext_env = - List.fold_left (fun env fid -> add_rel (Name fid,()) env) env lf in + List.fold_left (fun env fid -> Idset.add fid env) env lf in let defl = Array.of_list (List.map (dbrec ext_env) lt) in let arityl = Array.of_list (List.map (dbrec env) lA) in RRec (loc,RFix (Array.of_list ln,n), Array.of_list lf, arityl, defl) @@ -306,17 +310,17 @@ let dbize k sigma env allow_soapp lvar = with Not_found -> error_fixname_unbound "dbize (COFIX)" true locid iddef in let ext_env = - List.fold_left (fun env fid -> add_rel (Name fid,()) env) env lf in + List.fold_left (fun env fid -> Idset.add fid env) env lf in let defl = Array.of_list (List.map (dbrec ext_env) lt) in let arityl = Array.of_list (List.map (dbrec env) lA) in RRec (loc,RCofix n, Array.of_list lf, arityl, defl) | Node(loc,("PROD"|"LAMBDA" as k), [c1;Slam(_,ona,c2)]) -> - let na = match ona with - | Some s -> Name (id_of_string s) - | _ -> Anonymous in + let na,env' = match ona with + | Some s -> let id = id_of_string s in Name id, Idset.add id env + | _ -> Anonymous, env in let kind = if k="PROD" then BProd else BLambda in - RBinder(loc, kind, na, dbrec env c1, dbrec (add_rel (na,()) env) c2) + RBinder(loc, kind, na, dbrec env c1, dbrec env' c2) | Node(_,"PRODLIST", [c1;(Slam _ as c2)]) -> iterated_binder BProd 0 c1 env c2 @@ -326,11 +330,7 @@ let dbize k sigma env allow_soapp lvar = | Node(loc,"APPLISTEXPL", f::args) -> RApp (loc,dbrec env f,List.map (dbrec env) args) | Node(loc,"APPLIST", Nvar(locs,s)::args) -> - let (c, impargs) = - if List.mem s lvar then - RRef(loc,RVar (id_of_string s)),[] - else - dbize_ref k sigma env locs s + let (c, impargs) = dbize_ref k sigma env locs s lvar in RApp (loc, c, dbize_args env impargs args) | Node(loc,"APPLIST", f::args) -> @@ -393,18 +393,20 @@ let dbize k sigma env allow_soapp lvar = check_uppercase loc ids; check_number_of_pattern loc n pl; let env' = - List.fold_left (fun env id -> add_rel (Name id,()) env) env ids in + List.fold_left (fun env id -> Idset.add id env) env ids in (ids,pl,dbrec env' rhs) | _ -> anomaly "dbize: badly-formed ast for Cases equation" and iterated_binder oper n ty env = function | Slam(loc,ona,body) -> - let na = match ona with - | Some s -> check_capture s ty body; Name (id_of_string s) - | _ -> Anonymous + let na,env' = match ona with + | Some s -> + check_capture s ty body; + let id = id_of_string s in Name id, Idset.add id env + | _ -> Anonymous, env in RBinder(loc, oper, na, dbrec env ty, - (iterated_binder oper n ty (add_rel (na,()) env) body)) + (iterated_binder oper n ty env' body)) | body -> dbrec env body and dbize_args env l args = @@ -434,16 +436,23 @@ let dbize k sigma env allow_soapp lvar = * and translates a command to a constr. *) (*Takes a list of variables which must not be globalized*) +let from_list l = List.fold_right Idset.add l Idset.empty + let interp_rawconstr_gen sigma env allow_soapp lvar com = - dbize CCI sigma (unitize_env (context env)) allow_soapp lvar com + dbize CCI sigma + (from_list (ids_of_rel_context (rel_context env))) + allow_soapp (lvar,var_context env) com let interp_rawconstr sigma env com = interp_rawconstr_gen sigma env false [] com (*The same as interp_rawconstr but with a list of variables which must not be globalized*) + let interp_rawconstr_wo_glob sigma env lvar com = - dbize CCI sigma (unitize_env (context env)) false lvar com + dbize CCI sigma + (from_list (ids_of_rel_context (rel_context env))) + false (lvar,var_context env) com (*let raw_fconstr_of_com sigma env com = dbize_fw sigma (unitize_env (context env)) [] com @@ -460,7 +469,7 @@ let ast_adjust_consts sigma = (* locations are kept *) (let id = id_of_string s in if Ast.isMeta s then ast - else if List.mem id (ids_of_env env) then + else if Idset.mem id env then ast else try @@ -481,10 +490,10 @@ let ast_adjust_consts sigma = (* locations are kept *) with Not_found -> warning ("Could not globalize "^s); ast) - | Slam(loc,None,t) -> Slam(loc,None,dbrec (add_rel (Anonymous,()) env) t) + | Slam(loc,None,t) -> Slam(loc,None,dbrec env t) | Slam(loc,Some na,t) -> - let env' = add_rel (Name (id_of_string na),()) env in + let env' = Idset.add (id_of_string na) env in Slam(loc,Some na,dbrec env' t) | Node(loc,opn,tl) -> Node(loc,opn,List.map (dbrec env) tl) | x -> x @@ -493,7 +502,7 @@ let ast_adjust_consts sigma = (* locations are kept *) let globalize_command ast = let sign = Global.var_context () in - ast_adjust_consts Evd.empty (gLOB sign) ast + ast_adjust_consts Evd.empty (from_list (ids_of_var_context sign)) ast (* Avoid globalizing in non command ast for tactics *) let rec glob_ast sigma env = function @@ -502,16 +511,16 @@ let rec glob_ast sigma env = function | Node(loc,"COMMANDLIST",l) -> Node(loc,"COMMANDLIST", List.map (ast_adjust_consts sigma env) l) | Slam(loc,None,t) -> - Slam(loc,None,glob_ast sigma (add_rel (Anonymous,()) env) t) + Slam(loc,None,glob_ast sigma env t) | Slam(loc,Some na,t) -> - let env' = add_rel (Name (id_of_string na),()) env in + let env' = Idset.add (id_of_string na) env in Slam(loc,Some na, glob_ast sigma env' t) | Node(loc,opn,tl) -> Node(loc,opn,List.map (glob_ast sigma env) tl) | x -> x let globalize_ast ast = let sign = Global.var_context () in - glob_ast Evd.empty (gLOB sign) ast + glob_ast Evd.empty (from_list (ids_of_var_context sign)) ast (* Installation of the AST quotations. "command" is used by default. *) @@ -692,8 +701,10 @@ let pattern_of_rawconstr lvar c = let interp_constrpattern_gen sigma env lvar com = let c = - dbize CCI sigma (unitize_env (context env)) true (List.map (fun x -> - string_of_id (fst x)) lvar) com + dbize CCI sigma (from_list (ids_of_rel_context (rel_context env))) + true (List.map + (fun x -> + string_of_id (fst x)) lvar,var_context env) com and nlvar = List.map (fun (id,c) -> (id,pattern_of_constr c)) lvar in try pattern_of_rawconstr nlvar c diff --git a/parsing/pretty.ml b/parsing/pretty.ml index 9d189edfc..28c4cb3a9 100644 --- a/parsing/pretty.ml +++ b/parsing/pretty.ml @@ -42,9 +42,8 @@ let print_basename_mind sp mindid = let print_closed_sections = ref false let print_typed_value_in_env env (trm,typ) = - let sign = Environ.var_context env in - [< prterm_env (gLOB sign) trm ; 'fNL ; - 'sTR " : "; prtype_env (gLOB sign) typ ; 'fNL >] + [< prterm_env env trm ; 'fNL ; + 'sTR " : "; prtype_env env typ ; 'fNL >] let print_typed_value x = print_typed_value_in_env (Global.env ()) x @@ -80,30 +79,22 @@ let print_impl_args = function the pretty-print of a proposition (P:(nat->nat)->Prop)(P [u]u) synthesizes the type nat of the abstraction on u *) -let print_var name typ = +let print_var_def name body typ = + [< 'sTR "*** [" ; 'sTR name ; 'sPC; + hOV 0 [< 'sTR "="; prterm body; 'sPC; 'sTR ": "; prtype typ >] ; + 'sTR "]"; 'fNL >] + +let print_var_decl name typ = [< 'sTR "*** [" ; 'sTR name ; 'sTR " : "; prtype typ; 'sTR "]"; 'fNL >] -let print_env pk = - let pterminenv = (* if pk = FW then fprterm_env else *) prterm_env in - let pr_binder env (na,c) = - match na with - | Name id as name -> - let pc = pterminenv env c in [< print_id id; 'sTR":"; pc >] - | Anonymous -> anomaly "Anonymous variables in inductives" - in - let rec prec env = function - | [] -> [<>] - | [b] -> pr_binder env b - | b::rest -> - let pb = pr_binder env b in - let penvtl = prec (add_rel b env) rest in - [< pb; 'sTR";"; 'sPC; penvtl >] - in - prec (gLOB nil_sign) +let print_var (id,c,typ) = + let s = string_of_id id in + match c with + | Some body -> print_var_def s body typ + | None -> print_var_decl s typ let assumptions_for_print lna = - List.fold_right (fun na env -> add_rel (na,()) env) lna - (ENVIRON(nil_sign,nil_dbsign)) + List.fold_right (fun na env -> add_name na env) lna empty_names_context let implicit_args_id id l = if l = [] then @@ -126,26 +117,24 @@ let implicit_args_msg sp mipv = let print_mutual sp mib = let pk = kind_of_path sp in - let pterm,pterminenv = pkprinters pk in let env = Global.env () in let evd = Evd.empty in let {mind_packets=mipv; mind_nparams=nparams} = mib in let (lpars,_) = decomp_n_prod env evd nparams (mind_user_arity mipv.(0)) in - let lparsname = List.map fst lpars in - let lna = Array.map (fun mip -> Name mip.mind_typename) mipv in - let lparsprint = assumptions_for_print lparsname in - let prass ass_name (id,c) = - let pc = pterminenv ass_name c in [< print_id id; 'sTR " : "; pc >] + let arities = Array.map (fun mip -> (Name mip.mind_typename, None, mip.mind_nf_arity)) mipv in + let env_ar = push_rels lpars env in + let env_cstr = push_rels lpars (push_rels (Array.to_list arities) env) in + let pr_constructor (id,c) = + let pc = prterm_env env_cstr c in [< print_id id; 'sTR " : "; pc >] in let print_constructors mip = let lC = mind_user_lc mip in - let ass_name = assumptions_for_print (lparsname@(Array.to_list lna)) in let lidC = Array.to_list (array_map2 (fun id c -> (id,snd (decomp_n_prod env evd nparams c))) mip.mind_consnames lC) in let plidC = prlist_with_sep (fun () -> [<'bRK(1,0); 'sTR "| " >]) - (prass ass_name) lidC in + pr_constructor lidC in (hV 0 [< 'sTR " "; plidC >]) in let print_oneind mip = @@ -156,9 +145,8 @@ let print_mutual sp mib = if nparams = 0 then [<>] else - [< 'sTR "["; print_env pk (List.rev lpars); - 'sTR "]"; 'bRK(1,2) >]; - 'sTR ": "; pterminenv lparsprint arity; 'sTR " :=" >]); + [< 'sTR "["; pr_rel_context env lpars; 'sTR "]"; 'bRK(1,2) >]; + 'sTR ": "; prterm_env env_ar arity; 'sTR " :=" >]); 'bRK(1,2); print_constructors mip >]) in let mip = mipv.(0) in @@ -170,8 +158,8 @@ let print_mutual sp mib = if nparams = 0 then [<>] else - [< 'sTR" ["; print_env pk (List.rev lpars); 'sTR "]">]; - 'bRK(1,5); 'sTR": "; pterminenv lparsprint arity; 'sTR" :="; + [< 'sTR" ["; pr_rel_context env lpars; 'sTR "]">]; + 'bRK(1,5); 'sTR": "; prterm_env env_ar arity; 'sTR" :="; 'bRK(0,4); print_constructors mip; 'fNL; implicit_args_msg sp mipv >] ) (* Mutual [co]inductive definitions *) @@ -211,10 +199,10 @@ let print_extracted_mutual sp = print_mutual fwsp (Global.lookup_mind fwsp) | Some a -> fprterm a -let print_variable sp = - let (name,typ,_,_) = out_variable sp in - let l = implicits_of_var name in - [< print_var (string_of_id name) typ; print_impl_args l; 'fNL >] +let print_section_variable sp = + let ((id,_,_) as d,_,_) = out_variable sp in + let l = implicits_of_var id in + [< print_var d; print_impl_args l; 'fNL >] let print_constant with_values sep sp = let cb = Global.lookup_constant sp in @@ -251,7 +239,7 @@ let print_leaf_entry with_values sep (sp,lobj) = let tag = object_tag lobj in match (sp,tag) with | (_,"VARIABLE") -> - print_variable sp + print_section_variable sp | (_,("CONSTANT"|"PARAMETER")) -> print_constant with_values sep sp | (_,"INDUCTIVE") -> @@ -332,39 +320,37 @@ let list_filter_vec f vec = of the object, the assumptions that will make it possible to print its type, and the constr term that represent its type. *) -let print_constructors fn lna mip = - let ass_name = assumptions_for_print lna in +let print_constructors fn env_ar mip = let _ = - array_map2 (fun id c -> fn (string_of_id id) ass_name c) + array_map2 (fun id c -> fn (string_of_id id) env_ar c) mip.mind_consnames (mind_user_lc mip) in () -let crible (fn : string -> unit assumptions -> constr -> unit) name = - let hyps = gLOB (Global.var_context()) in +let crible (fn : string -> env -> constr -> unit) name = + let env = Global.env () in let imported = Library.opened_modules() in let const = global_reference CCI name in let rec crible_rec = function | (spopt,Lib.Leaf lobj)::rest -> (match (spopt,object_tag lobj) with | (_,"VARIABLE") -> - let (namec,typ,_,_) = out_variable spopt in + let ((idc,_,typ),_,_) = out_variable spopt in if (head_const (body_of_type typ)) = const then - fn (string_of_id namec) hyps (body_of_type typ); + fn (string_of_id idc) env (body_of_type typ); crible_rec rest | (sp,("CONSTANT"|"PARAMETER")) -> let {const_type=typ} = Global.lookup_constant sp in if (head_const (body_of_type typ)) = const then - fn (print_basename sp) hyps (body_of_type typ); + fn (print_basename sp) env (body_of_type typ); crible_rec rest | (sp,"INDUCTIVE") -> let mib = Global.lookup_mind sp in - let lna = - array_map_to_list - (fun mip -> Name mip.mind_typename) mib.mind_packets in + let arities = array_map_to_list (fun mip -> (Name mip.mind_typename, None, mip.mind_nf_arity)) mib.mind_packets in + let env_ar = push_rels arities env in (match const with | (DOPN(MutInd(sp',tyi),_)) -> if sp = objsp_of sp' then - print_constructors fn lna + print_constructors fn env_ar (mind_nth_type_packet mib tyi) | _ -> ()); crible_rec rest @@ -423,25 +409,20 @@ let print_name name = | _ -> raise Not_found in print_leaf_entry true " = " (sp,lobj) - with - | Not_found -> - (try - (match Declare.global_reference CCI name with - | VAR _ -> - let (_,typ) = Global.lookup_var name in - [< print_var str typ; - try print_impl_args (implicits_of_var name) - with _ -> [<>] >] - | DOPN(Const sp,_) -> - print_constant true " = " sp - | DOPN(MutInd (sp,_),_) -> - print_inductive sp - | DOPN (MutConstruct((sp,_),_),_) -> - print_inductive sp - | _ -> assert false) - with Not_found | Invalid_argument _ -> - error (str ^ " not a defined object")) - | Invalid_argument _ -> error (str ^ " not a defined object") + with Not_found -> + try + match fst (Declare.global_operator CCI name) with + | Const sp -> print_constant true " = " sp + | MutInd (sp,_) -> print_inductive sp + | MutConstruct((sp,_),_) -> print_inductive sp + | _ -> assert false + with Not_found -> + try + let (c,typ) = Global.lookup_var name in + [< print_var (name,c,typ); + try print_impl_args (implicits_of_var name) + with _ -> [<>] >] + with Not_found -> error (str ^ " not a defined object") let print_opaque_name name = let sigma = Evd.empty in @@ -463,8 +444,8 @@ let print_opaque_name name = let ty = Typeops.type_of_constructor env sigma cstr in print_typed_value (x, ty) | IsVar id -> - let a = snd(lookup_sign id sign) in - print_var (string_of_id id) a + let (c,ty) = lookup_var id env in + print_var (id,c,ty) | _ -> failwith "print_name" with Not_found -> error ((string_of_id name) ^ " not declared") @@ -475,9 +456,9 @@ let print_local_context () = | [] -> [< >] | (sp,Lib.Leaf lobj)::rest -> if "VARIABLE" = object_tag lobj then - let (name,typ,_,_) = out_variable sp in + let (d,_,_) = out_variable sp in [< print_var_rec rest; - print_var (string_of_id name) typ >] + print_var d >] else print_var_rec rest | _::rest -> print_var_rec rest diff --git a/parsing/pretty.mli b/parsing/pretty.mli index e10c53b80..dd7378d41 100644 --- a/parsing/pretty.mli +++ b/parsing/pretty.mli @@ -13,13 +13,12 @@ open Environ (* A Pretty-Printer for the Calculus of Inductive Constructions. *) -val assumptions_for_print : name list -> unit assumptions +val assumptions_for_print : name list -> names_context val print_basename : section_path -> string val print_basename_mind : section_path -> identifier -> string val print_closed_sections : bool ref val print_impl_args : int list -> std_ppcmds -val print_env : path_kind -> (name * constr) list -> std_ppcmds val print_context : bool -> Lib.library_segment -> std_ppcmds val print_library_entry : bool -> (section_path * Lib.node) -> std_ppcmds val print_full_context : unit -> std_ppcmds @@ -50,7 +49,7 @@ val print_coercions : unit -> std_ppcmds val print_path_between : identifier -> identifier -> std_ppcmds -val crible : (string -> unit assumptions -> constr -> unit) -> identifier -> +val crible : (string -> env -> constr -> unit) -> identifier -> unit val inspect : int -> std_ppcmds diff --git a/parsing/printer.ml b/parsing/printer.ml index 89bbdb2e9..da78f8352 100644 --- a/parsing/printer.ml +++ b/parsing/printer.ml @@ -86,28 +86,28 @@ let gentermpr gt = (* [at_top] means ids of env must be avoided in bound variables *) let gentermpr_core at_top env t = - gentermpr (Termast.ast_of_constr at_top (unitize_env env) t) + gentermpr (Termast.ast_of_constr at_top env t) let prterm_env_at_top env = gentermpr_core true env let prterm_env env = gentermpr_core false env -let prterm = prterm_env (gLOB nil_sign) +let prterm = prterm_env empty_env let prtype_env env typ = prterm_env env (incast_type typ) -let prtype = prtype_env (gLOB nil_sign) +let prtype = prtype_env empty_env let prjudge_env env j = (prterm_env env j.uj_val, prterm_env env (incast_type j.uj_type)) -let prjudge = prjudge_env (gLOB nil_sign) +let prjudge = prjudge_env empty_env (* Plus de "k"... let gentermpr k = gentermpr_core false let gentermpr_at_top k = gentermpr_core true let fprterm_env a = gentermpr FW a -let fprterm = fprterm_env (gLOB nil_sign) +let fprterm = fprterm_env empty_env let fprtype_env env typ = fprterm_env env (incast_type typ) -let fprtype = fprtype_env (gLOB nil_sign) +let fprtype = fprtype_env empty_env *) let fprterm_env a = @@ -123,27 +123,23 @@ let fprtype a = (* For compatibility *) let fterm0 = fprterm_env -let pr_constant env cst = gentermpr (ast_of_constant (unitize_env env) cst) -let pr_existential env ev = gentermpr (ast_of_existential (unitize_env env) ev) -let pr_inductive env ind = gentermpr (ast_of_inductive (unitize_env env) ind) +let pr_constant env cst = gentermpr (ast_of_constant env cst) +let pr_existential env ev = gentermpr (ast_of_existential env ev) +let pr_inductive env ind = gentermpr (ast_of_inductive env ind) let pr_constructor env cstr = - gentermpr (ast_of_constructor (unitize_env env) cstr) + gentermpr (ast_of_constructor env cstr) open Pattern let pr_ref_label = function (* On triche sur le contexte *) - | ConstNode sp -> pr_constant (Global.context()) (sp,[||]) - | IndNode sp -> pr_inductive (Global.context()) (sp,[||]) - | CstrNode sp -> pr_constructor (Global.context()) (sp,[||]) + | ConstNode sp -> pr_constant (Global.env()) (sp,[||]) + | IndNode sp -> pr_inductive (Global.env()) (sp,[||]) + | CstrNode sp -> pr_constructor (Global.env()) (sp,[||]) | VarNode id -> print_id id let pr_cases_pattern t = gentermpr (Termast.ast_of_cases_pattern t) let pr_rawterm t = gentermpr (Termast.ast_of_rawconstr t) -let pr_pattern_env env t = - gentermpr (Termast.ast_of_pattern (unitize_env env) t) -let pr_pattern t = pr_pattern_env (gLOB nil_sign) t - -(* For compatibility *) -let term0 = prterm_env +let pr_pattern_env env t = gentermpr (Termast.ast_of_pattern env t) +let pr_pattern t = pr_pattern_env empty_names_context t let rec gentacpr gt = Esyntax.genprint default_tacpr tactic_syntax_universe tactic_initial_prec gt @@ -155,86 +151,87 @@ and default_tacpr = function [< 'sTR s; 'bRK(1,2); hOV 0 (prlist_with_sep pr_spc gentacpr ta) >] | gt -> dfltpr gt -let print_decl sign (s,typ) = - let ptyp = prterm_env (gLOB sign) (body_of_type typ) in - [< print_id s ; 'sTR" : "; ptyp >] - -let print_binding env (na,typ) = - let ptyp = prterm_env env (body_of_type typ) in +let pr_var_decl env (id,c,typ) = + let pbody = match c with + | None -> [< >] + | Some c -> [< 'sTR" := "; prterm_env env c >] in + let ptyp = [< 'sTR" : "; prtype_env env typ >] in + [< print_id id ; hOV 0 [< pbody; ptyp >] >] + +let pr_rel_decl env (na,c,typ) = + let pbody = match c with + | None -> [< >] + | Some c -> [< 'sTR" :="; 'sPC; prterm_env env c >] in + let ptyp = prtype_env env typ in match na with - | Anonymous -> [< 'sTR"<>" ; 'sTR" : " ; ptyp >] - | Name id -> [< print_id id ; 'sTR" : "; ptyp >] + | Anonymous -> [< 'sTR"<>" ; 'sPC; pbody; 'sTR" :"; 'sPC; ptyp >] + | Name id -> [< print_id id ; 'sPC; pbody; 'sTR" :"; 'sPC; ptyp >] (* Prints out an "env" in a nice format. We print out the * signature,then a horizontal bar, then the debruijn environment. * It's printed out from outermost to innermost, so it's readable. *) -let sign_it_with f sign e = - snd (sign_it (fun id t (sign,e) -> (add_sign (id,t) sign, f id t sign e)) - sign (nil_sign,e)) - -let sign_it_with_i f sign e = - let (_,_,e') = - (sign_it (fun id t (i,sign,e) -> (i+1,add_sign (id,t) sign, - f i id t sign e)) - sign (0,nil_sign,e)) - in - e' - -let dbenv_it_with f env e = - snd (dbenv_it (fun na t (env,e) -> (add_rel (na,t) env, f na t env e)) - env (gLOB(get_globals env),e)) - (* Prints a signature, all declarations on the same line if possible *) -let pr_sign sign = - hV 0 [< (sign_it_with (fun id t sign pps -> - [< pps; 'wS 2; print_decl sign (id,t) >]) - sign) [< >] >] +let pr_var_context_of env = + hV 0 [< (fold_var_context + (fun env d pps -> [< pps; 'wS 2; pr_var_decl env d >]) + env) [< >] >] + +let pr_rel_context env rel_context = + let rec prec env = function + | [] -> [<>] + | [b] -> pr_rel_decl env b + | b::rest -> + let pb = pr_rel_decl env b in + let penvtl = prec (push_rel b env) rest in + [< pb; 'sTR";"; 'sPC; penvtl >] + in + prec env (List.rev rel_context) (* Prints an env (variables and de Bruijn). Separator: newline *) -let pr_env env = +let pr_context_unlimited env = let sign_env = - sign_it_with - (fun id t sign pps -> - let pidt = print_decl sign (id,t) in [< pps; 'fNL; pidt >]) - (get_globals env) [< >] + fold_var_context + (fun env d pps -> + let pidt = pr_var_decl env d in [< pps; 'fNL; pidt >]) + env [< >] in let db_env = - dbenv_it_with - (fun na t env pps -> - let pnat = print_binding env (na,t) in [< pps; 'fNL; pnat >]) + fold_rel_context + (fun env d pps -> + let pnat = pr_rel_decl env d in [< pps; 'fNL; pnat >]) env [< >] in [< sign_env; db_env >] -let pr_ne_env header k = function - | ENVIRON (sign,_) as env when isnull_sign sign & isnull_rel_env env -> [< >] - | env -> let penv = pr_env env in [< header; penv >] +let pr_ne_context_of header k env = + if Environ.context env = empty_context then [< >] + else let penv = pr_context_unlimited env in [< header; penv >] -let pr_env_limit n env = - let sign = get_globals env in - let lgsign = sign_length sign in +let pr_context_limit n env = + let var_context = Environ.var_context env in + let lgsign = List.length var_context in if n >= lgsign then - pr_env env + pr_context_unlimited env else let k = lgsign-n in - let sign_env = - sign_it_with_i - (fun i id t sign pps -> + let _,sign_env = + fold_var_context + (fun env d (i,pps) -> if i < k then - [< pps ;'sTR "." >] + (i+1, [< pps ;'sTR "." >]) else - let pidt = print_decl sign (id,t) in - [< pps ; 'fNL ; - 'sTR (emacs_str (String.make 1 (Char.chr 253))); - pidt >]) - (get_globals env) [< >] + let pidt = pr_var_decl env d in + (i+1, [< pps ; 'fNL ; + 'sTR (emacs_str (String.make 1 (Char.chr 253))); + pidt >])) + env (0,[< >]) in let db_env = - dbenv_it_with - (fun na t env pps -> - let pnat = print_binding env (na,t) in + fold_rel_context + (fun env d pps -> + let pnat = pr_rel_decl env d in [< pps; 'fNL; 'sTR (emacs_str (String.make 1 (Char.chr 253))); pnat >]) @@ -242,6 +239,6 @@ let pr_env_limit n env = in [< sign_env; db_env >] -let pr_env_opt env = match Options.print_hyps_limit () with - | None -> hV 0 (pr_env env) - | Some n -> hV 0 (pr_env_limit n env) +let pr_context_of env = match Options.print_hyps_limit () with + | None -> hV 0 (pr_context_unlimited env) + | Some n -> hV 0 (pr_context_limit n env) diff --git a/parsing/printer.mli b/parsing/printer.mli index ff3179c7d..67434ee31 100644 --- a/parsing/printer.mli +++ b/parsing/printer.mli @@ -6,6 +6,7 @@ open Pp open Names open Term open Sign +open Environ open Rawterm open Pattern (*i*) @@ -13,40 +14,43 @@ open Pattern (* These are the entry points for printing terms, context, tac, ... *) val gentacpr : Coqast.t -> std_ppcmds -val prterm_env : 'a assumptions -> constr -> std_ppcmds -val prterm_env_at_top : 'a assumptions -> constr -> std_ppcmds +val prterm_env : env -> constr -> std_ppcmds +val prterm_env_at_top : env -> constr -> std_ppcmds val prterm : constr -> std_ppcmds -val prtype_env : 'a assumptions -> typed_type -> std_ppcmds +val prtype_env : env -> typed_type -> std_ppcmds val prtype : typed_type -> std_ppcmds val prjudge_env : - 'a assumptions -> Environ.unsafe_judgment -> std_ppcmds * std_ppcmds + env -> Environ.unsafe_judgment -> std_ppcmds * std_ppcmds val prjudge : Environ.unsafe_judgment -> std_ppcmds * std_ppcmds val pr_rawterm : Rawterm.rawconstr -> std_ppcmds val pr_cases_pattern : Rawterm.cases_pattern -> std_ppcmds -val pr_constant : 'a assumptions -> constant -> std_ppcmds -val pr_existential : 'a assumptions -> existential -> std_ppcmds -val pr_constructor : 'a assumptions -> constructor -> std_ppcmds -val pr_inductive : 'a assumptions -> inductive -> std_ppcmds +val pr_constant : env -> constant -> std_ppcmds +val pr_existential : env -> existential -> std_ppcmds +val pr_constructor : env -> constructor -> std_ppcmds +val pr_inductive : env -> inductive -> std_ppcmds val pr_ref_label : constr_label -> std_ppcmds val pr_pattern : constr_pattern -> std_ppcmds -val pr_pattern_env : 'a assumptions -> constr_pattern -> std_ppcmds +val pr_pattern_env : names_context -> constr_pattern -> std_ppcmds -val pr_ne_env : std_ppcmds -> path_kind -> context -> std_ppcmds +val pr_ne_context_of : std_ppcmds -> path_kind -> env -> std_ppcmds -val pr_sign : var_context -> std_ppcmds -val pr_env_opt : context -> std_ppcmds +val pr_var_decl : env -> var_declaration -> std_ppcmds +val pr_rel_decl : env -> rel_declaration -> std_ppcmds -val emacs_str : string -> string +val pr_var_context_of : env -> std_ppcmds +val pr_rel_context : env -> rel_context -> std_ppcmds +val pr_context_of : env -> std_ppcmds -(* For compatibility *) -val term0 : 'a assumptions -> constr -> std_ppcmds +val emacs_str : string -> string -val fprterm_env : 'a assumptions -> constr -> std_ppcmds +(*i*) +val fprterm_env : env -> constr -> std_ppcmds val fprterm : constr -> std_ppcmds -val fprtype_env : 'a assumptions -> typed_type -> std_ppcmds +val fprtype_env : env -> typed_type -> std_ppcmds val fprtype : typed_type -> std_ppcmds (* For compatibility *) -val fterm0 : 'a assumptions -> constr -> std_ppcmds +val fterm0 : env -> constr -> std_ppcmds +(*i*) diff --git a/parsing/termast.ml b/parsing/termast.ml index 7c4568af1..ff8222e4e 100644 --- a/parsing/termast.ml +++ b/parsing/termast.ml @@ -9,6 +9,7 @@ open Generic open Term open Inductive open Sign +open Environ open Declare open Impargs open Coqast @@ -349,7 +350,7 @@ let occur_id env id0 c = | DLAMV(_,v) -> array_exists (occur (n+1)) v | Rel p -> p>n & - (try (fst (lookup_rel (p-n) env) = Name id0) + (try lookup_name_of_rel (p-n) env = Name id0 with Not_found -> false) (* Unbound indice : may happen in debug *) | DOP0 _ -> false in @@ -396,8 +397,6 @@ let global_vars_and_consts t = list_uniquize (collect [] t) let used_of = global_vars_and_consts -let ids_of_env = Sign.ids_of_env - (****************************************************************************) (* Tools for printing of Cases *) @@ -490,14 +489,14 @@ let ids_of_var cl = (* Main translation function from constr -> ast *) let old_bdize at_top env t = - let init_avoid = if at_top then ids_of_env env else [] in + let init_avoid = if at_top then ids_of_context env else [] in let rec bdrec avoid env t = match collapse_appl t with (* Not well-formed constructions *) | DLAM(na,c) -> (match concrete_name true (*On ne sait pas si prod*)avoid env na c with | (Some id,avoid') -> slam(Some (string_of_id id), - bdrec avoid' (add_rel (Name id,()) env) c) + bdrec avoid' (add_name (Name id) env) c) | (None,avoid') -> slam(None,bdrec avoid' env (pop c))) | DLAMV(na,cl) -> @@ -508,14 +507,14 @@ let old_bdize at_top env t = let id = next_name_away na avoid in slam(Some (string_of_id id), ope("V$", array_map_to_list - (bdrec (id::avoid) (add_rel (Name id,()) env)) cl)) + (bdrec (id::avoid) (add_name (Name id) env)) cl)) (* Well-formed constructions *) | regular_constr -> (match kind_of_term regular_constr with | IsRel n -> (try - match fst(lookup_rel n env) with + match lookup_name_of_rel n env with | Name id -> nvar (string_of_id id) | Anonymous -> raise Not_found with Not_found -> @@ -610,7 +609,7 @@ let old_bdize at_top env t = let lfi = List.map (fun na -> next_name_away na avoid) lfn in let def_env = List.fold_left - (fun env id -> add_rel (Name id,()) env) env lfi in + (fun env id -> add_name (Name id) env) env lfi in let def_avoid = lfi@avoid in let f = List.nth lfi n in let rec split_lambda binds env avoid = function @@ -621,7 +620,7 @@ let old_bdize at_top env t = let id = next_name_away na avoid in let ast_of_bind = ope("BINDER",[ast;nvar (string_of_id id)]) in - let new_env = add_rel (Name id,()) env in + let new_env = add_name (Name id) env in split_lambda (ast_of_bind::binds) new_env (id::avoid) (n-1,b) | _ -> error "split_lambda" @@ -632,7 +631,7 @@ let old_bdize at_top env t = | (n, DOP2(Prod,t,DLAM(na,b))) -> let ast = bdrec avoid env t in let id = next_name_away na avoid in - let new_env = add_rel (Name id,()) env in + let new_env = add_name (Name id) env in split_product new_env (id::avoid) (n-1,b) | _ -> error "split_product" in @@ -655,7 +654,7 @@ let old_bdize at_top env t = let lfi = List.map (fun na -> next_name_away na avoid) lfn in let def_env = List.fold_left - (fun env id -> add_rel (Name id,()) env) env lfi in + (fun env id -> add_name (Name id) env) env lfi in let def_avoid = lfi@avoid in let f = List.nth lfi n in let listdecl = @@ -684,7 +683,7 @@ let old_bdize at_top env t = if not print_underscore or (dependent (Rel 1) b) then x else Anonymous in let id = next_name_away x' avoid in - let new_env = (add_rel (Name id,()) env) in + let new_env = (add_name (Name id) env) in let new_avoid = id::avoid in let new_nvarlist = (nvar (string_of_id id))::nvarlist in buildrec new_nvarlist new_avoid new_env (n-1,b) @@ -697,7 +696,7 @@ let old_bdize at_top env t = -> (* nommage de la nouvelle variable *) let id = next_ident_away (id_of_string "x") avoid in let new_nvarlist = (nvar (string_of_id id))::nvarlist in - let new_env = (add_rel (Name id,()) env) in + let new_env = (add_name (Name id) env) in let new_avoid = id::avoid in let new_b = DOPN(AppL,[|lift 1 b; Rel 1|]) in buildrec new_nvarlist new_avoid new_env (n-1,new_b) @@ -707,8 +706,8 @@ let old_bdize at_top env t = and factorize_binder n avoid env oper na ty c = let (env2, avoid2,na2) = match concrete_name (oper=Lambda) avoid env na c with - (Some id,l') -> add_rel (Name id,()) env, l', Some (string_of_id id) - | (None,l') -> add_rel (Anonymous,()) env, l', None + (Some id,l') -> add_name (Name id) env, l', Some (string_of_id id) + | (None,l') -> add_name Anonymous env, l', None in let (p,body) = match c with DOP2(oper',ty',DLAM(na',c')) @@ -720,7 +719,7 @@ let old_bdize at_top env t = in (p,slam(na2, body)) in - bdrec init_avoid env t + bdrec init_avoid (names_of_rel_context env) t (* FIN TO EJECT *) (******************************************************************) @@ -728,10 +727,11 @@ let ast_of_constr at_top env t = let t' = if !print_casts then t else Reduction.strong (fun _ _ -> strip_outer_cast) - Environ.empty_env Evd.empty t in + empty_env Evd.empty t in try - let avoid = if at_top then ids_of_env env else [] in - ast_of_raw (Detyping.detype avoid env t') + let avoid = if at_top then ids_of_context env else [] in + ast_of_raw + (Detyping.detype avoid (names_of_rel_context env) t') with Detyping.StillDLAM -> old_bdize at_top env t' @@ -744,14 +744,15 @@ let ast_of_inductive env = ast_of_inductive_ref (ast_of_constr false env) let ast_of_constructor env = ast_of_constructor_ref (ast_of_constr false env) let rec ast_of_pattern env = function - | PRef ref -> ast_of_ref (ast_of_constr false env) ref + | PRef ref -> ast_of_ref (fun c -> ast_of_raw (Detyping.detype [] env c)) ref + | PRel n -> - (try match fst (lookup_rel n env) with + (try match lookup_name_of_rel n env with | Name id -> ast_of_ident id | Anonymous -> anomaly "ast_of_pattern: index to an anonymous variable" with Not_found -> - let s = "[REL "^(string_of_int (number_of_rels env - n))^"]" + let s = "[REL "^(string_of_int n)^"]" in nvar s) | PApp (f,args) -> @@ -771,7 +772,7 @@ let rec ast_of_pattern env = function | PBinder (BProd,Anonymous,t,c) -> ope("PROD",[ast_of_pattern env t; slam(None,ast_of_pattern env c)]) | PBinder (bk,na,t,c) -> - let env' = add_rel (na,()) env in + let env' = add_name na env in let (n,a) = factorize_binder_pattern env' 1 bk na (ast_of_pattern env t) c in let tag = match bk with @@ -785,7 +786,7 @@ let rec ast_of_pattern env = function ope(tag,[ast_of_pattern env t;a]) | PLet (id,a,t,c) -> - let c' = ast_of_pattern (add_rel (Name id,()) env) c in + let c' = ast_of_pattern (add_name (Name id) env) c in ope("LET",[ast_of_pattern env a; slam(Some (string_of_id id),c')]) @@ -814,7 +815,7 @@ and factorize_binder_pattern env n oper na aty c = when (oper = oper') & (aty = ast_of_pattern env ty') & not (na' = Anonymous & oper = BProd) -> - factorize_binder_pattern (add_rel (na',()) env) (n+1) oper na' aty c' + factorize_binder_pattern (add_name na' env) (n+1) oper na' aty c' | _ -> (n,ast_of_pattern env c) in (p,slam(stringopt_of_name na, body)) diff --git a/parsing/termast.mli b/parsing/termast.mli index 7961ee7fb..4e62177a9 100644 --- a/parsing/termast.mli +++ b/parsing/termast.mli @@ -5,6 +5,7 @@ open Names open Term open Sign +open Environ open Rawterm open Pattern (*i*) @@ -14,21 +15,21 @@ open Pattern val ast_of_cases_pattern : cases_pattern -> Coqast.t val ast_of_rawconstr : rawconstr -> Coqast.t -val ast_of_pattern : unit assumptions -> constr_pattern -> Coqast.t +val ast_of_pattern : names_context -> constr_pattern -> Coqast.t (* If [b=true] in [ast_of_constr b env c] then the variables in the first level of quantification clashing with the variables in [env] are renamed *) -val ast_of_constr : bool -> unit assumptions -> constr -> Coqast.t +val ast_of_constr : bool -> env -> constr -> Coqast.t (*i C'est bdize qui sait maintenant s'il faut afficher les casts ou pas val bdize_no_casts : bool -> unit assumptions -> constr -> Coqast.t i*) -val ast_of_constant : unit assumptions -> constant -> Coqast.t -val ast_of_existential : unit assumptions -> existential -> Coqast.t -val ast_of_constructor : unit assumptions -> constructor -> Coqast.t -val ast_of_inductive : unit assumptions -> inductive -> Coqast.t +val ast_of_constant : env -> constant -> Coqast.t +val ast_of_existential : env -> existential -> Coqast.t +val ast_of_constructor : env -> constructor -> Coqast.t +val ast_of_inductive : env -> inductive -> Coqast.t (* For debugging *) val print_implicits : bool ref diff --git a/pretyping/cases.ml b/pretyping/cases.ml index 47fcdcc15..1d25ba84d 100644 --- a/pretyping/cases.ml +++ b/pretyping/cases.ml @@ -25,20 +25,20 @@ let mkExistential isevars env = new_isevar isevars env (mkCast dummy_sort dummy_sort) CCI let norec_branch_scheme env isevars cstr = - prod_it (mkExistential isevars env) cstr.cs_args + it_mkProd_or_LetIn (mkExistential isevars env) cstr.cs_args -let lift_args l = list_map_i (fun i (name,c) -> (name,liftn 1 i c)) 1 l - -let rec_branch_scheme env isevars ((sp,j),_) recargs cstr = +let rec_branch_scheme env isevars ((sp,j),_) recargs cstr = let rec crec (args,recargs) = match args, recargs with - | (name,c)::rea,(ra::reca) -> - DOP2(Prod,c,DLAM(name, + | (name,None,c)::rea,(ra::reca) -> + DOP2(Prod,body_of_type c,DLAM(name, match ra with | Mrec k when k=j -> mkArrow (mkExistential isevars env) - (crec (lift_args rea,reca)) + (crec (List.rev (lift_rel_context 1 (List.rev rea)),reca)) | _ -> crec (rea,reca))) + | (name,Some d,c)::rea, reca -> failwith "TODO" +(* mkLetIn (name,d,body_of_type c,crec (rea,reca))) *) | [],[] -> mkExistential isevars env | _ -> anomaly "rec_branch_scheme" in @@ -145,7 +145,7 @@ let inductive_of_rawconstructor c = (* Environment management *) let push_rel_type sigma (na,t) env = - push_rel (na,get_assumption_of env sigma t) env + push_rel_decl (na,get_assumption_of env sigma t) env let push_rels vars env = List.fold_right (fun nvar env -> push_rel_type Evd.empty nvar env) vars env @@ -537,6 +537,39 @@ let prepare_unif_pb typ cs = (* This is the problem: finding P s.t. cs_args |- (P realargs ci) = p' *) (Array.map (lift (-n)) cs.cs_concl_realargs, ci, p') +(* +(* Infering the predicate *) +let prepare_unif_pb typ cs = + let n = cs.cs_nargs in + let _,p = decompose_prod_n n typ in + let ci = build_dependent_constructor cs in + (* This is the problem: finding P s.t. cs_args |- (P realargs ci) = p *) + (n, cs.cs_concl_realargs, ci, p) + +let eq_operator_lift k (n,n') = function + | OpRel p, OpRel p' when p > k & p' > k -> + if p < k+n or p' < k+n' then false else p - n = p' - n' + | op, op' -> op = op' + +let rec transpose_args n = + if n=0 then [] + else + (Array.map (fun l -> List.hd l) lv):: + (transpose_args (m-1) (Array.init (fun l -> List.tl l))) + +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) !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 + let argvl = transpose_args (List.length args1) pv' in + let k' = shift_operator k op1 in + let argl = List.map (unify_clauses k') argvl in + gather_constr (reloc_operator (k,n1) op1) argl +*) let abstract_conclusion typ cs = let n = cs.cs_nargs in @@ -545,18 +578,17 @@ let abstract_conclusion typ cs = let infer_predicate env isevars typs cstrs (IndFamily (mis,_) as indf) = (* Il faudra substituer les isevars a un certain moment *) - 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 (sign,_) = get_arity indf in - if array_for_all (fun (_,_,typ) -> the_conv_x env isevars typn typ) eqns - then - let pred = lam_it (lift (List.length sign) typn) sign in - (false,pred) (* true = dependent -- par défaut *) + if Array.length cstrs = 0 then + failwith "TODO4-3" else - if Array.length typs = 0 then - failwith "TODO4-3" + 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 (sign,_) = get_arity indf in + if array_for_all (fun (_,_,typ) -> the_conv_x env isevars typn typ) eqns + then + let pred = lam_it (lift (List.length sign) typn) sign in + (false,pred) (* true = dependent -- par défaut *) else let s = get_sort_of env !isevars typs.(0) in let predpred = lam_it (mkSort s) sign in @@ -717,7 +749,8 @@ let shift_problem pb = (* Building the sub-pattern-matching problem for a given branch *) let build_branch pb defaults current eqns const_info = - let names = get_names pb.env const_info.cs_args eqns in + let cs_args = decls_of_rel_context const_info.cs_args in + let names = get_names pb.env cs_args eqns in let newpats = if !substitute_alias then List.map (fun na -> PatVar (dummy_loc,na)) names @@ -726,7 +759,7 @@ let build_branch pb defaults current eqns const_info = let submatdef = List.map (expand_defaults newpats current) defaults in let submat = List.map (fun (tms,eqn) -> prepend_pattern tms eqn) eqns in if submat = [] & submatdef = [] then error "Non exhaustive"; - let typs = List.map2 (fun (_,t) na -> (na,t)) const_info.cs_args (List.rev names) in + let typs = List.map2 (fun (_,t) na -> (na,t)) cs_args (List.rev names) in let tomatchs = List.fold_left2 (fun l (na,t) dep_in_rhs -> @@ -834,7 +867,7 @@ let matx_of_eqns env eqns = let build_eqn (ids,lpatt,rhs) = let rhs = { rhs_env = env; - other_ids = ids@(ids_of_sign (get_globals (context env))); + other_ids = ids@(ids_of_var_context (var_context env)); private_ids = []; user_ids = ids; subst = []; @@ -874,7 +907,7 @@ let inh_coerce_to_ind isevars env ty tyi = List.fold_right (fun (na,ty) (env,evl) -> let aty = get_assumption_of env Evd.empty ty in - (push_rel (na,aty) env, + (push_rel_decl (na,aty) env, (new_isevar isevars env (incast_type aty) CCI)::evl)) ntys (env,[]) in let expected_typ = applist (mkMutInd tyi,evarl) in diff --git a/pretyping/coercion.ml b/pretyping/coercion.ml index 1e90ef8f3..8269a4741 100644 --- a/pretyping/coercion.ml +++ b/pretyping/coercion.ml @@ -130,7 +130,7 @@ let rec inh_conv_coerce_to_fail env isevars c1 hj = (* let jv1 = exemeta_rec def_vty_con env isevars v1 in let assv1 = assumption_of_judgement env !isevars jv1 in *) let assv1 = outcast_type v1 in - let env1 = push_rel (x,assv1) env in + let env1 = push_rel_decl (x,assv1) env in let h2 = inh_conv_coerce_to_fail env isevars u2 {uj_val = v2; uj_type = get_assumption_of env1 !isevars t2 } in @@ -140,7 +140,7 @@ let rec inh_conv_coerce_to_fail env isevars c1 hj = let name = (match name with | Anonymous -> Name (id_of_string "x") | _ -> name) in - let env1 = push_rel (name,assu1) env in + let env1 = push_rel_decl (name,assu1) env in let h1 = inh_conv_coerce_to_fail env isevars t1 {uj_val = Rel 1; diff --git a/pretyping/detyping.ml b/pretyping/detyping.ml index 01ed9dc3b..90697191a 100644 --- a/pretyping/detyping.ml +++ b/pretyping/detyping.ml @@ -30,7 +30,7 @@ open Rawterm type used_idents = identifier list -let occur_id env id0 c = +let occur_id env_names id0 c = let rec occur n = function | VAR id -> id=id0 | DOPN (Const sp,cl) -> (basename sp = id0) or (array_exists (occur n) cl) @@ -49,26 +49,27 @@ let occur_id env id0 c = | DLAMV(_,v) -> array_exists (occur (n+1)) v | Rel p -> p>n & - (try (fst (lookup_rel (p-n) env) = Name id0) - with Not_found -> false) (* Unbound indice : may happen in debug *) + (try lookup_name_of_rel (p-n) env_names = Name id0 + with Not_found -> false) (* Unbound indice: may happen in debug *) | DOP0 _ -> false in occur 1 c -let next_name_not_occuring name l env t = +let next_name_not_occuring name l env_names t = let rec next id = - if List.mem id l or occur_id env id t then next (lift_ident id) else id + if List.mem id l or occur_id env_names id t then next (lift_ident id) + else id in match name with | Name id -> next id | Anonymous -> id_of_string "_" (* Remark: Anonymous var may be dependent in Evar's contexts *) -let concrete_name l env n c = +let concrete_name l env_names n c = if n = Anonymous & not (dependent (Rel 1) c) then (None,l) else - let fresh_id = next_name_not_occuring n l env c in + let fresh_id = next_name_not_occuring n l env_names c in let idopt = if dependent (Rel 1) c then (Some fresh_id) else None in (idopt, fresh_id::l) @@ -95,8 +96,6 @@ let global_vars_and_consts t = list_uniquize (collect [] t) let used_of = global_vars_and_consts -let ids_of_env = Sign.ids_of_env - (****************************************************************************) (* Tools for printing of Cases *) @@ -262,22 +261,22 @@ let ids_of_var cl = (Array.to_list cl) *) -let lookup_name_as_renamed env t s = - let rec lookup avoid env n = function - DOP2(Prod,_,DLAM(name,c')) -> - (match concrete_name avoid env name c' with +let lookup_name_as_renamed ctxt t s = + let rec lookup avoid env_names n = function + DOP2(Prod,t,DLAM(name,c')) -> + (match concrete_name avoid env_names name c' with (Some id,avoid') -> if id=s then (Some n) - else lookup avoid' (add_rel (Name id,()) env) (n+1) c' - | (None,avoid') -> lookup avoid' env (n+1) (pop c')) - | DOP2(Cast,c,_) -> lookup avoid env n c + else lookup avoid' (add_name (Name id) env_names) (n+1) c' + | (None,avoid') -> lookup avoid' env_names (n+1) (pop c')) + | DOP2(Cast,c,_) -> lookup avoid env_names n c | _ -> None - in lookup (ids_of_env env) env 1 t + in lookup (ids_of_var_context ctxt) empty_names_context 1 t let lookup_index_as_renamed t n = let rec lookup n d = function DOP2(Prod,_,DLAM(name,c')) -> - (match concrete_name [] (gLOB nil_sign) name c' with + (match concrete_name [] empty_names_context name c' with (Some _,_) -> lookup n (d+1) c' | (None ,_) -> if n=1 then Some d else lookup (n-1) (d+1) c') | DOP2(Cast,c,_) -> lookup n d c @@ -294,11 +293,11 @@ let rec detype avoid env t = | regular_constr -> (match kind_of_term regular_constr with | IsRel n -> - (try match fst (lookup_rel n env) with + (try match lookup_name_of_rel n env with | Name id -> RRef (dummy_loc, RVar id) | Anonymous -> anomaly "detype: index to an anonymous variable" with Not_found -> - let s = "[REL "^(string_of_int (number_of_rels env - n))^"]" + let s = "[REL "^(string_of_int n)^"]" in RRef (dummy_loc, RVar (id_of_string s))) | IsMeta n -> RMeta (dummy_loc, n) | IsVar id -> RRef (dummy_loc,RVar id) @@ -363,7 +362,7 @@ and detype_fix fk avoid env cl lfn vt = let lfi = List.map (fun id -> next_name_away id avoid) lfn in let def_avoid = lfi@avoid in let def_env = - List.fold_left (fun env id -> add_rel (Name id,()) env) env lfi in + List.fold_left (fun env id -> add_name (Name id) env) env lfi in RRec(dummy_loc,fk,Array.of_list lfi,Array.map (detype avoid env) cl, Array.map (detype def_avoid def_env) vt) @@ -372,9 +371,9 @@ and detype_eqn avoid env constr_id construct_nargs branch = let make_pat x avoid env b ids = if not (force_wildcard ()) or (dependent (Rel 1) b) then let id = next_name_away_with_default "x" x avoid in - PatVar (dummy_loc,Name id),id::avoid,(add_rel (Name id,()) env),id::ids + PatVar (dummy_loc,Name id),id::avoid,(add_name (Name id) env),id::ids else - PatVar (dummy_loc,Anonymous),avoid,(add_rel (Anonymous,()) env),ids + PatVar (dummy_loc,Anonymous),avoid,(add_name Anonymous env),ids in let rec buildrec ids patlist avoid env = function | 0 , rhs -> @@ -405,4 +404,4 @@ and detype_binder bk avoid env na ty c = | (None,l') -> Anonymous, l' in RBinder (dummy_loc,bk, na',detype [] env ty, - detype avoid' (add_rel (na',()) env) c) + detype avoid' (add_name na' env) c) diff --git a/pretyping/detyping.mli b/pretyping/detyping.mli index 9e0a30d09..e961cfe91 100644 --- a/pretyping/detyping.mli +++ b/pretyping/detyping.mli @@ -5,6 +5,7 @@ open Names open Term open Sign +open Environ open Rawterm (*i*) @@ -13,11 +14,10 @@ open Rawterm exception StillDLAM -val detype : identifier list -> unit assumptions -> constr -> rawconstr +val detype : identifier list -> names_context -> constr -> rawconstr (* look for the index of a named var or a nondep var as it is renamed *) -val lookup_name_as_renamed : - unit assumptions -> constr -> identifier -> int option +val lookup_name_as_renamed : var_context -> constr -> identifier -> int option val lookup_index_as_renamed : constr -> int -> int option diff --git a/pretyping/evarconv.ml b/pretyping/evarconv.ml index a4faeb3e4..3afc73469 100644 --- a/pretyping/evarconv.ml +++ b/pretyping/evarconv.ml @@ -261,9 +261,9 @@ and evar_eqappr_x env isevars pbty appr1 appr2 = | ((DOP2(Prod,c1,DLAM(n,c2)),[]), (DOP2(Prod,c'1,DLAM(_,c'2)),[])) -> evar_conv_x env isevars CONV c1 c'1 - & evar_conv_x - (push_rel (n,Retyping.get_assumption_of env !isevars (nf_ise1 !isevars c1)) env) isevars - pbty c2 c'2 + & + (let d = Retyping.get_assumption_of env !isevars (nf_ise1 !isevars c1) + in evar_conv_x (push_rel_decl (n,d) env) isevars pbty c2 c'2) | ((DOPN(MutInd _ as o1,cl1) as ind1,l'1), (DOPN(MutInd _ as o2,cl2) as ind2,l'2)) -> diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml index 2de1219e3..804d635db 100644 --- a/pretyping/evarutil.ml +++ b/pretyping/evarutil.ml @@ -29,6 +29,7 @@ let distinct_id_list = in drec [] +(* let filter_sign p sign x = sign_it (fun id ty (v,ids,sgn) -> @@ -36,21 +37,21 @@ let filter_sign p sign x = if disc then (v', id::ids, sgn) else (v', ids, add_sign (id,ty) sgn)) sign (x,[],nil_sign) - +*) (*------------------------------------* * functional operations on evar sets * *------------------------------------*) (* All ids of sign must be distincts! *) -let new_isevar_sign env sigma typ args = +let new_isevar_sign env sigma typ instance = let sign = var_context env in - if not (list_distinct (ids_of_sign sign)) then + if not (list_distinct (ids_of_var_context sign)) then error "new_isevar_sign: two vars have the same name"; let newev = Evd.new_evar() in let info = { evar_concl = typ; evar_env = env; evar_body = Evar_empty; evar_info = None } in - (Evd.add sigma newev info, mkEvar newev args) + (Evd.add sigma newev info, mkEvar (newev,Array.of_list instance)) (* 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... *) @@ -60,8 +61,8 @@ let dummy_sort = mkType dummy_univ cumulativity now includes Prop and Set in Type. *) let new_type_var env sigma = let sign = var_context env in - let args = (Array.of_list (List.map mkVar (ids_of_sign sign))) in - let (sigma',c) = new_isevar_sign env sigma dummy_sort args in + let instance = List.map mkVar (ids_of_var_context sign) in + let (sigma',c) = new_isevar_sign env sigma dummy_sort instance in (sigma', c) let split_evar_to_arrow sigma c = @@ -70,16 +71,16 @@ let split_evar_to_arrow sigma c = let evenv = evd.evar_env in let (sigma1,dom) = new_type_var evenv sigma in let hyps = var_context evenv in - let nvar = next_ident_away (id_of_string "x") (ids_of_sign hyps) in - let nevenv = push_var (nvar,make_typed dom (Type dummy_univ)) evenv in - let (sigma2,rng) = new_type_var nevenv sigma1 in - let prod = mkProd (named_hd nevenv dom Anonymous) dom (subst_var nvar rng) in + let nvar = next_ident_away (id_of_string "x") (ids_of_var_context hyps) in + let newenv = push_var_decl (nvar,make_typed dom (Type dummy_univ)) evenv in + let (sigma2,rng) = new_type_var newenv sigma1 in + let prod = mkProd (named_hd newenv dom Anonymous) dom (subst_var nvar rng) in let sigma3 = Evd.define sigma2 ev prod in let dsp = num_of_evar dom in let rsp = num_of_evar rng in (sigma3, - mkEvar dsp args, - mkEvar rsp (array_cons (mkRel 1) (Array.map (lift 1) args))) + mkEvar (dsp,args), + mkEvar (rsp, array_cons (mkRel 1) (Array.map (lift 1) args))) (* Redefines an evar with a smaller context (i.e. it may depend on less @@ -99,17 +100,17 @@ let do_restrict_hyps sigma c = let (_,(rsign,ncargs)) = List.fold_left (fun (sign,(rs,na)) a -> - (tl_sign sign, + (List.tl sign, if not(closed0 a) then (rs,na) else - (add_sign (hd_sign sign) rs, a::na))) - (hyps,(nil_sign,[])) args + (add_var (List.hd sign) rs, a::na))) + (hyps,([],[])) args in - let sign' = rev_sign rsign in + let sign' = List.rev rsign in let env' = change_hyps (fun _ -> sign') env in - let args' = Array.of_list (List.map mkVar (ids_of_sign sign')) in - let (sigma',nc) = new_isevar_sign env' sigma evd.evar_concl args' in + let instance = List.map mkVar (ids_of_var_context sign') in + let (sigma',nc) = new_isevar_sign env' sigma evd.evar_concl instance in let sigma'' = Evd.define sigma' ev nc in (sigma'', nc) @@ -199,29 +200,16 @@ let real_clean isevars sp args rhs = (* [new_isevar] declares a new existential in an env env with type typ *) (* Converting the env into the sign of the evar to define *) -let append_rels_to_vars ctxt = - dbenv_it - (fun na t (subst,sign) -> - let na = if na = Anonymous then Name(id_of_string"_") else na in - let id = next_name_away na (ids_of_sign sign) in - ((VAR id)::subst, add_sign (id,typed_app (substl subst) t) sign)) - ctxt ([],get_globals ctxt) - let new_isevar isevars env typ k = - let ctxt = context env in - let subst,sign = append_rels_to_vars ctxt in + let subst,env' = push_rels_to_vars env in let typ' = substl subst typ in - let env' = change_hyps (fun _ -> sign) env in - let newargs = - (List.rev(rel_list 0 (number_of_rels ctxt))) - @(List.map (fun id -> VAR id) (ids_of_sign (get_globals ctxt))) in - let (sigma',evar) = - new_isevar_sign env' !isevars typ' (Array.of_list newargs) in + let instance = + (List.rev (rel_list 0 (rel_context_length (rel_context env)))) + @(List.map mkVar (ids_of_var_context (var_context env))) in + let (sigma',evar) = new_isevar_sign env' !isevars typ' instance in isevars := sigma'; evar - - (* [evar_define] solves the problem lhs = rhs when lhs is an uninstantiated * evar, i.e. tries to find the body ?sp for lhs=DOPN(Const sp,args) * ?sp [ sp.hyps \ args ] unifies to rhs @@ -247,7 +235,7 @@ let evar_define isevars lhs rhs = let evd = ise_map isevars ev in let hyps = var_context evd.evar_env in (* the substitution to invert *) - let worklist = List.combine (ids_of_sign hyps) args in + let worklist = List.combine (ids_of_var_context hyps) args in let body = real_clean isevars ev worklist rhs in ise_define isevars ev body; [ev] @@ -269,19 +257,19 @@ let solve_refl conv_algo isevars c1 c2 = array_fold_left2 (fun (sgn,rsgn) a1 a2 -> if conv_algo a1 a2 then - (tl_sign sgn, add_sign (hd_sign sgn) rsgn) + (List.tl sgn, add_var (List.hd sgn) rsgn) else - (tl_sign sgn, rsgn)) - (hyps,nil_sign) argsv1 argsv2 + (List.tl sgn, rsgn)) + (hyps,[]) argsv1 argsv2 in - let nsign = rev_sign rsign in + let nsign = List.rev rsign in let nenv = change_hyps (fun _ -> nsign) env in - let nargs = (Array.of_list (List.map mkVar (ids_of_sign nsign))) in + let nargs = (Array.of_list (List.map mkVar (ids_of_var_context nsign))) in let newev = Evd.new_evar () in let info = { evar_concl = evd.evar_concl; evar_env = nenv; evar_body = Evar_empty; evar_info = None } in isevars := - Evd.define (Evd.add !isevars newev info) ev (mkEvar newev nargs); + Evd.define (Evd.add !isevars newev info) ev (mkEvar (newev,nargs)); Some [ev] @@ -419,8 +407,7 @@ let split_tycon loc env isevars = function isevars := sigma; Some dom, Some rng | _ -> - Stdpp.raise_with_loc loc (Type_errors.TypeError (CCI,context env,Type_errors.NotProduct c))) + Stdpp.raise_with_loc loc (Type_errors.TypeError (CCI,env,Type_errors.NotProduct c))) let valcon_of_tycon x = x - diff --git a/pretyping/evarutil.mli b/pretyping/evarutil.mli index f648e4a83..dbebeac98 100644 --- a/pretyping/evarutil.mli +++ b/pretyping/evarutil.mli @@ -17,10 +17,6 @@ open Reduction val filter_unique : 'a list -> 'a list val distinct_id_list : identifier list -> identifier list -val filter_sign : - ('a -> identifier * typed_type -> bool * 'a) -> var_context -> 'a -> - 'a * identifier list * var_context - val dummy_sort : constr val do_restrict_hyps : 'a evar_map -> constr -> 'a evar_map * constr val has_ise : 'a evar_map -> constr -> bool diff --git a/pretyping/pretype_errors.ml b/pretyping/pretype_errors.ml index 9f711f390..247620a67 100644 --- a/pretyping/pretype_errors.ml +++ b/pretyping/pretype_errors.ml @@ -11,51 +11,51 @@ let raise_pretype_error (loc,k,ctx,te) = Stdpp.raise_with_loc loc (TypeError(k,ctx,te)) let error_var_not_found_loc loc k s = - raise_pretype_error (loc,k, Global.context() (*bidon*), VarNotFound s) + raise_pretype_error (loc,k, Global.env() (*bidon*), VarNotFound s) let error_cant_find_case_type_loc loc env expr = - raise_pretype_error (loc, CCI, context env, CantFindCaseType expr) + raise_pretype_error (loc, CCI, env, CantFindCaseType expr) let error_actual_type_loc loc env c actty expty = - raise_pretype_error (loc, CCI, context env, ActualType (c,actty,expty)) + raise_pretype_error (loc, CCI, env, ActualType (c,actty,expty)) let error_cant_apply_not_functional_loc loc env rator randl = raise_pretype_error - (loc,CCI,context env, CantApplyNonFunctional (rator,randl)) + (loc,CCI,env, CantApplyNonFunctional (rator,randl)) let error_cant_apply_bad_type_loc loc env t rator randl = - raise_pretype_error (loc, CCI, context env, CantApplyBadType (t,rator,randl)) + raise_pretype_error (loc, CCI, env, CantApplyBadType (t,rator,randl)) let error_ill_formed_branch k env c i actty expty = - raise (TypeError (k, context env, IllFormedBranch (c,i,actty,expty))) + raise (TypeError (k, env, IllFormedBranch (c,i,actty,expty))) let error_number_branches_loc loc k env c ct expn = - raise_pretype_error (loc, k, context env, NumberBranches (c,ct,expn)) + raise_pretype_error (loc, k, env, NumberBranches (c,ct,expn)) let error_case_not_inductive_loc loc k env c ct = - raise_pretype_error (loc, k, context env, CaseNotInductive (c,ct)) + raise_pretype_error (loc, k, env, CaseNotInductive (c,ct)) (* Pattern-matching errors *) let error_bad_constructor_loc loc k cstr ind = - raise_pretype_error (loc, k, Global.context(), BadConstructor (cstr,ind)) + raise_pretype_error (loc, k, Global.env(), BadConstructor (cstr,ind)) let error_wrong_numarg_constructor_loc loc k c n = - raise_pretype_error (loc, k, Global.context(), WrongNumargConstructor (c,n)) + raise_pretype_error (loc, k, Global.env(), WrongNumargConstructor (c,n)) let error_wrong_predicate_arity_loc loc k env c n1 n2 = - raise_pretype_error (loc, k, context env, WrongPredicateArity (c,n1,n2)) + raise_pretype_error (loc, k, env, WrongPredicateArity (c,n1,n2)) let error_needs_inversion k env x t = - raise (TypeError (k, context env, NeedsInversion (x,t))) + raise (TypeError (k, env, NeedsInversion (x,t))) let error_ill_formed_branch_loc loc k env c i actty expty = - raise_pretype_error (loc, k, context env, IllFormedBranch (c,i,actty,expty)) + raise_pretype_error (loc, k, env, IllFormedBranch (c,i,actty,expty)) let error_occur_check k env ev c = - raise (TypeError (k, context env, OccurCheck (ev,c))) + raise (TypeError (k, env, OccurCheck (ev,c))) let error_not_clean k env ev c = - raise (TypeError (k, context env, NotClean (ev,c))) + raise (TypeError (k, env, NotClean (ev,c))) diff --git a/pretyping/pretyping.ml b/pretyping/pretyping.ml index 68a75ee32..cadaa7f78 100644 --- a/pretyping/pretyping.ml +++ b/pretyping/pretyping.ml @@ -189,21 +189,19 @@ let evar_type_case isevars env ct pt lft p c = in check_branches_message isevars env (c,ct) (bty,lft); (mind,rslty) *) -let pretype_var loc env lvar id = +let pretype_id loc env lvar id = try List.assoc id lvar - with - Not_found -> - (try - match lookup_id id (context env) with - | RELNAME (n,typ) -> - { uj_val = Rel n; - uj_type = typed_app (lift n) typ } - | GLOBNAME (id,typ) -> - { uj_val = VAR id; - uj_type = typ } - with Not_found -> - error_var_not_found_loc loc CCI id);; + with Not_found -> + try + let (n,typ) = lookup_rel_id id (rel_context env) in + { uj_val = Rel n; uj_type = typed_app (lift n) typ } + with Not_found -> + try + let typ = lookup_id_type id (var_context env) in + { uj_val = VAR id; uj_type = typ } + with Not_found -> + error_var_not_found_loc loc CCI id (*************************************************************************) (* Main pretyping function *) @@ -212,7 +210,7 @@ let trad_metamap = ref [] let trad_nocheck = ref false let pretype_ref pretype loc isevars env lvar = function -| RVar id -> pretype_var loc env lvar id +| RVar id -> pretype_id loc env lvar id | RConst (sp,ctxt) -> let cst = (sp,Array.map pretype ctxt) in @@ -287,6 +285,16 @@ match cstr with (* Où teste-t-on que le résultat doit satisfaire tycon ? *) | RHole loc -> if !compter then nbimpl:=!nbimpl+1; +(* OLD + (match vtcon with + (true,(Some v, _)) -> + {uj_val=v.utj_val; uj_type=make_typed (mkSort v.utj_type) (Type Univ.dummy_univ)} + | (false,(None,Some ty)) -> + let c = new_isevar isevars env ty CCI in + {uj_val=c;uj_type=make_typed ty (Type Univ.dummy_univ)} + | (true,(None,None)) -> + let ty = mkCast dummy_sort dummy_sort in +*) (match tycon with | Some ty -> let c = new_isevar isevars env ty CCI in @@ -305,7 +313,7 @@ match cstr with (* Où teste-t-on que le résultat doit satisfaire tycon ? *) let nbfix = Array.length lfi in let lfi = List.map (fun id -> Name id) (Array.to_list lfi) in let newenv = - array_fold_left2 (fun env id ar -> (push_rel (id,ar) env)) + array_fold_left2 (fun env id ar -> (push_rel_decl (id,ar) env)) env (Array.of_list (List.rev lfi)) (vect_lift_type lara) in let vdefj = Array.mapi @@ -345,7 +353,7 @@ match cstr with (* Où teste-t-on que le résultat doit satisfaire tycon ? *) let j = pretype_type dom_valcon env isevars lvar lmeta c1 in let assum = assumption_of_type_judgment (inh_ass_of_j env isevars j) in let var = (name,assum) in - let j' = pretype rng (push_rel var env) isevars lvar lmeta c2 + let j' = pretype rng (push_rel_decl var env) isevars lvar lmeta c2 in fst (abs_rel env !isevars name assum j') @@ -353,7 +361,7 @@ match cstr with (* Où teste-t-on que le résultat doit satisfaire tycon ? *) let j = pretype empty_tycon env isevars lvar lmeta c1 in let assum = inh_ass_of_j env isevars j in let var = (name,assumption_of_type_judgment assum) in - let j' = pretype empty_tycon (push_rel var env) isevars lvar lmeta c2 in + let j' = pretype empty_tycon (push_rel_decl var env) isevars lvar lmeta c2 in (try fst (gen_rel env !isevars name assum j') with TypeError _ as e -> Stdpp.raise_with_loc loc e) diff --git a/pretyping/retyping.ml b/pretyping/retyping.ml index 60647a946..72a3d3d51 100644 --- a/pretyping/retyping.ml +++ b/pretyping/retyping.ml @@ -43,7 +43,7 @@ let rec type_of env cstr= IsMeta n -> (try strip_outer_cast (List.assoc n metamap) with Not_found -> anomaly "type_of: this is not a well-typed term") - | IsRel n -> lift n (body_of_type (snd (lookup_rel n env))) + | IsRel n -> lift n (body_of_type (snd (lookup_rel_type n env))) | IsVar id -> (try body_of_type (snd (lookup_var id env)) with Not_found -> @@ -65,7 +65,7 @@ let rec type_of env cstr= whd_betadeltaiota env sigma (applist (p,al)) | IsLambda (name,c1,c2) -> let var = make_typed c1 (sort_of env c1) in - mkProd name c1 (type_of (push_rel (name,var) env) c2) + mkProd name c1 (type_of (push_rel_decl (name,var) env) c2) | IsFix ((vn,i),(lar,lfi,vdef)) -> lar.(i) | IsCoFix (i,(lar,lfi,vdef)) -> lar.(i) | IsAppL(f,args)-> @@ -81,7 +81,7 @@ and sort_of env t = | IsSort (Type u) -> Type Univ.dummy_univ | IsProd (name,t,c2) -> let var = make_typed t (sort_of env t) in - (match (sort_of (push_rel (name,var) env) c2) with + (match (sort_of (push_rel_decl (name,var) env) c2) with | Prop _ as s -> s | Type u2 -> Type Univ.dummy_univ) | IsAppL(f,args) -> sort_of_atomic_type env sigma (type_of env f) args diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index f97b32cd1..a0fe84ee6 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -26,69 +26,59 @@ let rev_firstn_liftn fn ln = in rfprec fn [] -let make_elim_fun f largs = - try - let (sp, _) = destConst f in - match constant_eval sp with - | EliminationCases n when List.length largs >= n -> f - | EliminationFix (lv,n) when List.length largs >= n -> - let labs,_ = list_chop n largs in - let p = List.length lv in - let la' = list_map_i - (fun q aq -> - try (Rel (p+1-(list_index (n+1-q) (List.map fst lv)))) - with Not_found -> aq) 1 - (List.map (lift p) labs) - in - list_fold_left_i - (fun i c (k,a) -> - DOP2(Lambda,(substl (rev_firstn_liftn (n-k) (-i) la') a), - DLAM(Name(id_of_string"x"),c))) 0 (applistc f la') lv - | _ -> raise Redelimination - with Failure _ -> - raise Redelimination - -(* F is convertible to DOPN(Fix(recindices,bodynum),bodyvect) make - the reduction using this extra information *) +(* 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 ... *) -let rebuild_global_name id = - let sp = Nametab.sp_of_id CCI id in - let (oper,hyps) = Declare.global_operator sp id in - DOPN(oper,Array.of_list(List.map (fun id -> VAR id) (Sign.ids_of_sign hyps))) +let make_elim_fun f lv n largs = + let (sp,args) = destConst f in + let labs,_ = list_chop n largs in + let p = List.length lv in + let ylv = List.map fst lv in + let la' = list_map_i + (fun q aq -> + try (Rel (p+1-(list_index (n-q) ylv))) + with Not_found -> aq) 0 + (List.map (lift p) labs) + in + fun id -> + let fi = DOPN(Const(make_path (dirpath sp) id (kind_of_path sp)),args) in + list_fold_left_i + (fun i c (k,a) -> + DOP2(Lambda,(substl (rev_firstn_liftn (n-k) (-i) la') a), + DLAM(Name(id_of_string"x"),c))) 0 (applistc fi la') lv -let mydestFix = function - | DOPN (Fix (recindxs,i),a) -> - let (types,funnames,bodies) = destGralFix a in - (recindxs,i,funnames,bodies) - | _ -> invalid_arg "destFix" +(* [f] is convertible to [DOPN(Fix(recindices,bodynum),bodyvect)] make + the reduction using this extra information *) -let contract_fix_use_function f fix = - let (recindices,bodynum,names,bodies) = mydestFix fix in +let contract_fix_use_function f + ((recindices,bodynum),(types,names,bodies as typedbodies)) = let nbodies = Array.length recindices in let make_Fi j = - let id = match List.nth names j with Name id -> id | _ -> assert false in - rebuild_global_name id in - let lbodies = list_assign (list_tabulate make_Fi nbodies) bodynum f in - substl (List.rev lbodies) bodies.(bodynum) + match List.nth names j with Name id -> f id | _ -> assert false in + let lbodies = list_tabulate make_Fi nbodies in + substl (List.rev lbodies) bodies.(bodynum) let reduce_fix_use_function f whfun fix stack = - match fix with - | DOPN (Fix(recindices,bodynum),bodyvect) -> - (match fix_recarg fix stack with - | None -> (false,(fix,stack)) - | Some (recargnum,recarg) -> - let (recarg'hd,_ as recarg')= whfun recarg [] in - let stack' = list_assign stack recargnum (applist recarg') in - (match recarg'hd with - | DOPN(MutConstruct _,_) -> - (true,(contract_fix_use_function f fix,stack')) - | _ -> (false,(fix,stack')))) - | _ -> assert false + let dfix = destFix fix in + match fix_recarg dfix stack with + | None -> (false,(fix,stack)) + | Some (recargnum,recarg) -> + let (recarg'hd,_ as recarg')= whfun recarg [] in + let stack' = list_assign stack recargnum (applist recarg') in + (match recarg'hd with + | DOPN(MutConstruct _,_) -> + (true,(contract_fix_use_function f dfix,stack')) + | _ -> (false,(fix,stack'))) -let contract_cofix_use_function f (bodynum,(_,_,bodies as typedbodies)) = +let contract_cofix_use_function f (bodynum,(_,names,bodies as typedbodies)) = let nbodies = Array.length bodies in - let make_Fi j = mkCoFix (nbodies-j-1,typedbodies) in - let subbodies = list_assign (list_tabulate make_Fi nbodies) bodynum f in + let make_Fi j = + match List.nth names j with Name id -> f id | _ -> assert false in + let subbodies = list_tabulate make_Fi nbodies in substl subbodies bodies.(bodynum) let reduce_mind_case_use_function env f mia = @@ -106,12 +96,14 @@ let special_red_case env whfun p c ci lf = let rec redrec c l = let (constr,cargs) = whfun c l in match constr with - | DOPN(Const _,_) as g -> + | DOPN(Const sp,args) as g -> if evaluable_constant env g then let gvalue = constant_value env g in if reducible_mind_case gvalue then - reduce_mind_case_use_function env g - {mP=p; mconstr=gvalue; mcargs=cargs; mci=ci; mlf=lf} + let build_fix_name id = + DOPN(Const(make_path (dirpath sp) id (kind_of_path sp)),args) + in reduce_mind_case_use_function env build_fix_name + {mP=p; mconstr=gvalue; mcargs=cargs; mci=ci; mlf=lf} else redrec gvalue cargs else @@ -127,17 +119,30 @@ let special_red_case env whfun p c ci lf = let rec red_elim_const env sigma k largs = if not (evaluable_constant env k) then raise Redelimination; - let f = make_elim_fun k largs in - match whd_betadeltaeta_stack env sigma (constant_value env k) largs with - | (DOPN(MutCase _,_) as mc,lrest) -> - let (ci,p,c,lf) = destCase mc in - (special_red_case env (construct_const env sigma) p c ci lf, lrest) - | (DOPN(Fix _,_) as fix,lrest) -> - let (b,(c,rest)) = - reduce_fix_use_function f (construct_const env sigma) fix lrest - in - if b then (nf_beta env sigma c, rest) else raise Redelimination - | _ -> assert false + let (sp, args) = destConst k in + let elim_style = constant_eval sp in + match elim_style with + | EliminationCases n when List.length largs >= n -> begin + let c = constant_value env k in + match whd_betadeltaeta_stack env sigma c largs with + | (DOPN(MutCase _,_) as mc,lrest) -> + let (ci,p,c,lf) = destCase mc in + (special_red_case env (construct_const env sigma) p c ci lf, + lrest) + | _ -> assert false + end + | EliminationFix (lv,n) when List.length largs >= n -> begin + let c = constant_value env k in + match whd_betadeltaeta_stack env sigma c largs with + | (DOPN(Fix _,_) as fix,lrest) -> + let f id = make_elim_fun k lv n largs id in + let (b,(c,rest)) = + reduce_fix_use_function f (construct_const env sigma) fix lrest + in + if b then (nf_beta env sigma c, rest) else raise Redelimination + | _ -> assert false + end + | _ -> raise Redelimination and construct_const env sigma c stack = let rec hnfstack x stack = diff --git a/pretyping/typing.ml b/pretyping/typing.ml index 985b1d399..0dc0d0a5b 100644 --- a/pretyping/typing.ml +++ b/pretyping/typing.ml @@ -92,7 +92,7 @@ let rec execute mf env sigma cstr = | IsLambda (name,c1,c2) -> let j = execute mf env sigma c1 in let var = assumption_of_judgment env sigma j in - let env1 = push_rel (name,var) env in + let env1 = push_rel_decl (name,var) env in let j' = execute mf env1 sigma c2 in let (j,_) = abs_rel env1 sigma name var j' in j @@ -101,7 +101,7 @@ let rec execute mf env sigma cstr = let j = execute mf env sigma c1 in let varj = type_judgment env sigma j in let var = assumption_of_type_judgment varj in - let env1 = push_rel (name,var) env in + let env1 = push_rel_decl (name,var) env in let j' = execute mf env1 sigma c2 in let (j,_) = gen_rel env1 sigma name varj j' in j @@ -119,7 +119,7 @@ and execute_fix mf env sigma lar lfi vdef = let nlara = List.combine (List.rev lfi) (Array.to_list (vect_lift_type lara)) in let env1 = - List.fold_left (fun env nvar -> push_rel nvar env) env nlara in + List.fold_left (fun env nvar -> push_rel_decl nvar env) env nlara in let vdefj = execute_array mf env1 sigma vdef in let vdefv = Array.map j_val_only vdefj in let cst3 = type_fixpoint env1 sigma lfi lara vdefj in diff --git a/proofs/clenv.ml b/proofs/clenv.ml index 827618856..4c2cbf987 100644 --- a/proofs/clenv.ml +++ b/proofs/clenv.ml @@ -33,10 +33,9 @@ type 'a clausenv = { type wc = walking_constraints let new_evar_in_sign env = - let hyps = Environ.var_context env in + let ids = ids_of_var_context (Environ.var_context env) in let ev = new_evar () in - DOPN(Evar ev, - Array.of_list (List.map (fun id -> VAR id) (ids_of_sign hyps))) + DOPN(Evar ev, Array.of_list (List.map (fun id -> VAR id) ids)) let rec whd_evar env sigma = function | DOPN(Evar ev,_) as k -> diff --git a/proofs/evar_refiner.ml b/proofs/evar_refiner.ml index b73d19ee9..abfe2305d 100644 --- a/proofs/evar_refiner.ml +++ b/proofs/evar_refiner.ml @@ -100,7 +100,7 @@ let w_add_sign (id,t) (wc : walking_constraints) = ids_mk (ts_mod (fun evr -> { focus = evr.focus; - env = push_var (id,t) evr.env; + env = push_var_decl (id,t) evr.env; decls = evr.decls }) (ids_it wc)) diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index fa69b6e73..6bea1e819 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -116,11 +116,11 @@ let get_current_proof_name () = let add_proof (na,pfs,ts) = Edit.create proof_edits (na,pfs,ts,Some !undo_limit) -let del_proof na = +let delete_proof na = try Edit.delete proof_edits na with (UserError ("Edit.delete",_)) -> - errorlabstrm "Pfedit.del_proof" + errorlabstrm "Pfedit.delete_proof" [< 'sTR"No such proof" ; msg_proofs false >] let init_proofs () = Edit.clear proof_edits @@ -142,7 +142,7 @@ let restart_proof () = errorlabstrm "Pfedit.restart" [< 'sTR"No focused proof to restart" ; msg_proofs true >] | Some(na,_,ts) -> - del_proof na; + delete_proof na; start (na,ts); set_current_proof na @@ -178,16 +178,15 @@ let undo n = errorlabstrm "Pfedit.undo" [< 'sTR"No focused proof"; msg_proofs true >] (*********************************************************************) -(* Proof releasing *) +(* Proof cooking *) (*********************************************************************) -let release_proof () = +let cook_proof () = let (pfs,ts) = get_state() and ident = get_current_proof_name () in let {evar_concl=concl} = ts.top_goal and strength = ts.top_strength in let pfterm = extract_pftreestate pfs in - del_proof ident; (ident, ({ const_entry_body = Cooked pfterm; const_entry_type = Some concl }, strength)) @@ -204,11 +203,9 @@ let check_no_pending_proofs () = [< 'sTR"Proof editing in progress" ; (msg_proofs false) ; 'sTR"Use \"Abort All\" first or complete proof(s)." >] -let abort_proof = del_proof +let delete_current_proof () = delete_proof (get_current_proof_name ()) -let abort_current_proof () = del_proof (get_current_proof_name ()) - -let abort_all_proofs = init_proofs +let delete_all_proofs = init_proofs (*********************************************************************) (* Modifying the current prooftree *) @@ -238,14 +235,6 @@ let solve_nth k tac = let pft = get_pftreestate() in if not (List.mem (-1) (cursor_of_pftreestate pft)) then mutate (solve_nth_pftreestate k tac) - (* ; - let pfs =get_pftreestate() in - let pf = proof_of_pftreestate pfs in - let (pfterm,metamap) = refiner__extract_open_proof pf.goal.hyps pf - in (try typing__control_only_guard empty_evd pfterm - with e -> (undo 1; raise e)); - ()) - *) else error "cannot apply a tactic when we are descended behind a tactic-node" diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index 0bf91510b..a796cd1f4 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -29,19 +29,19 @@ val refining : unit -> bool val check_no_pending_proofs : unit -> unit -(*s [abort_proof name] aborts proof of name [name] or failed if no proof -has this name *) +(*s [delete_proof name] deletes proof of name [name] or failed if no proof + has this name *) -val abort_proof : identifier -> unit +val delete_proof : identifier -> unit -(* [abort_current_proof ()] aborts current focused proof or failed if +(* [delete_current_proof ()] deletes current focused proof or failed if no proof is focused *) -val abort_current_proof : unit -> unit +val delete_current_proof : unit -> unit -(* [abort_all_proofs ()] aborts all open proofs if any *) +(* [delete_all_proofs ()] deletes all open proofs if any *) -val abort_all_proofs : unit -> unit +val delete_all_proofs : unit -> unit (*s [undo n] undoes the effect of the last [n] tactics applied to the current proof; it fails if no proof is focused or if the ``undo'' @@ -82,13 +82,11 @@ val resume_proof : identifier -> unit val suspend_proof : unit -> unit -(*s [release_proof ()] turns the current proof into a constant with - its name and strength, then remove the proof from the set of - pending proofs; it fails if there is no current proof of if it is - not completed *) +(*s [cook_proof ()] turns the current proof (assumed completed) + into a constant with its name and strength; it fails if there is + no current proof of if it is not completed *) -val release_proof : unit -> - identifier * (Declarations.constant_entry * strength) +val cook_proof : unit -> identifier * (Declarations.constant_entry * strength) (*s [get_pftreestate ()] returns the current focused pending proof or raises [UserError "no focused proof"] *) @@ -145,22 +143,3 @@ val traverse : int -> unit val traverse_nth_goal : int -> unit val traverse_next_unproven : unit -> unit val traverse_prev_unproven : unit -> unit - -(*i N'ont pas à être exportées -type proof_topstate -val msg_proofs : bool -> std_ppcmds -val undo_limit : int ref -val get_state : unit -> pftreestate * proof_topstate -val get_topstate : unit -> proof_topstate -val add_proof : string * pftreestate * proof_topstate -> unit -val del_proof : string -> unit -val init_proofs : unit -> unit - -val mutate : (pftreestate -> pftreestate) -> unit -val rev_mutate : (pftreestate -> pftreestate) -> unit -val start : string * proof_topstate -> unit -val proof_term : unit -> constr -val save_anonymous : bool -> string -> unit -val traverse_to : int list -> unit -i*) - diff --git a/proofs/proof_trees.ml b/proofs/proof_trees.ml index e6150fba6..438f0efe3 100644 --- a/proofs/proof_trees.ml +++ b/proofs/proof_trees.ml @@ -147,7 +147,7 @@ let ctxt_access sigma sp = let pf_lookup_name_as_renamed hyps ccl s = - Detyping.lookup_name_as_renamed (gLOB hyps) ccl s + Detyping.lookup_name_as_renamed hyps ccl s let pf_lookup_index_as_renamed ccl n = Detyping.lookup_index_as_renamed ccl n @@ -166,7 +166,7 @@ open Printer term_env *) let pf_lookup_name_as_renamed hyps ccl s = - Detyping.lookup_name_as_renamed (gLOB hyps) ccl s + Detyping.lookup_name_as_renamed hyps ccl s let pf_lookup_index_as_renamed ccl n = Detyping.lookup_index_as_renamed ccl n @@ -174,16 +174,15 @@ let pf_lookup_index_as_renamed ccl n = let pr_idl idl = prlist_with_sep pr_spc print_id idl let pr_goal g = - let sign = context g.evar_env in - let penv = pr_env_opt sign in - let pc = prterm_env_at_top sign g.evar_concl in + let penv = pr_context_of g.evar_env in + let pc = prterm_env_at_top g.evar_env g.evar_concl in [< 'sTR" "; hV 0 [< penv; 'fNL; 'sTR (emacs_str (String.make 1 (Char.chr 253))) ; 'sTR "============================"; 'fNL ; 'sTR" " ; pc>]; 'fNL>] let pr_concl n g = - let pc = prterm_env_at_top (context g.evar_env) g.evar_concl in + let pc = prterm_env_at_top g.evar_env g.evar_concl in [< 'sTR (emacs_str (String.make 1 (Char.chr 253))) ; 'sTR "subgoal ";'iNT n;'sTR " is:";'cUT;'sTR" " ; pc >] @@ -233,16 +232,9 @@ let pr_seq evd = | Some i -> i | None -> anomaly "pr_seq : info = None" in - let hyps = var_context env in - let sign = List.rev (list_of_sign hyps) in let pc = pr_ctxt info in - let pdcl = - prlist_with_sep pr_spc - (fun (id,ty) -> - hOV 0 [< print_id id; 'sTR" : ";prterm (body_of_type ty) >]) - sign - in - let pcl = prterm_env_at_top (gLOB hyps) cl in + let pdcl = pr_var_context_of env in + let pcl = prterm_env_at_top env cl in hOV 0 [< pc; pdcl ; 'sPC ; hOV 0 [< 'sTR"|- " ; pcl >] >] let prgl gl = @@ -252,13 +244,13 @@ let prgl gl = let pr_evgl gl = let plc = pr_idl (List.map id_of_existential (lc_toList (get_lc gl))) in - let phyps = pr_idl (ids_of_sign (var_context gl.evar_env)) in + let phyps = pr_idl (ids_of_var_context (var_context gl.evar_env)) in let pc = prterm gl.evar_concl in hOV 0 [< 'sTR"[[" ; plc; 'sTR"] " ; phyps; 'sPC; 'sTR"|- " ; pc; 'sTR"]" >] let pr_evgl_sign gl = let plc = pr_idl (List.map id_of_existential (lc_toList (get_lc gl))) in - let ps = pr_sign (var_context gl.evar_env) in + let ps = pr_var_context_of gl.evar_env in let pc = prterm gl.evar_concl in hOV 0 [< 'sTR"[[" ; plc ; 'sTR"] " ; ps; 'sPC; 'sTR"|- " ; pc; 'sTR"]" >] @@ -352,7 +344,7 @@ let ast_of_cvt_arg = function | Integer n -> num n | Command c -> ope ("COMMAND",[c]) | Constr c -> - ope ("COMMAND",[ast_of_constr false (assumptions_for_print []) c]) + ope ("COMMAND",[ast_of_constr false (Global.env ()) c]) | Clause idl -> ope ("CLAUSE", List.map (compose nvar string_of_id) idl) | Bindings bl -> ope ("BINDINGS", List.map (ast_of_cvt_bind (fun x -> x)) bl) @@ -360,7 +352,7 @@ let ast_of_cvt_arg = function ope ("BINDINGS", List.map (ast_of_cvt_bind - (ast_of_constr false (assumptions_for_print []))) bl) + (ast_of_constr false (Global.env ()))) bl) | Tacexp ast -> ope ("TACTIC",[ast]) | Tac tac -> failwith "TODO: ast_of_cvt_arg: Tac" | Redexp red -> failwith "TODO: ast_of_cvt_arg: Redexp" diff --git a/proofs/refiner.ml b/proofs/refiner.ml index 1a16f7744..333287b3e 100644 --- a/proofs/refiner.ml +++ b/proofs/refiner.ml @@ -44,7 +44,7 @@ let norm_goal sigma gl = let red_fun = nf_ise1 sigma in let ncl = red_fun gl.evar_concl in { evar_concl = ncl; - evar_env = change_hyps (map_sign_typ (typed_app red_fun)) gl.evar_env; + evar_env = map_context red_fun gl.evar_env; evar_body = gl.evar_body; evar_info = gl.evar_info } @@ -74,9 +74,10 @@ let rec mapshape nl (fl : (proof_tree list -> proof_tree) list) let m,l = list_chop h l in (List.hd fl m) :: (mapshape t (List.tl fl) l) -(* given a proof p, frontier p gives (l,v) where l is the list of goals +(* frontier : proof_tree -> goal list * validation + given a proof p, frontier p gives (l,v) where l is the list of goals to be solved to complete the proof, and v is the corresponding validation *) - + let rec frontier p = match p.ref with | None -> @@ -230,17 +231,14 @@ let rc_of_pfsigma sigma = rc_of_gc sigma.sigma sigma.it.goal (* rc_of_glsigma : proof sigma -> readable_constraints *) let rc_of_glsigma sigma = rc_of_gc sigma.sigma sigma.it -(* extract_open_proof : constr signature -> proof - -> constr * (int * constr) list - takes a constr signature corresponding to global definitions - and a (not necessarly complete) proof - and gives a pair (pfterm,obl) +(* extract_open_proof : proof_tree -> constr * (int * constr) list + takes a (not necessarly complete) proof and gives a pair (pfterm,obl) where pfterm is the constr corresponding to the proof and obl is an int*constr list [ (m1,c1) ; ... ; (mn,cn)] where the mi are metavariables numbers, and ci are their types. Their proof should be completed in order to complete the initial proof *) -let extract_open_proof sign pf = +let extract_open_proof sigma pf = let open_obligations = ref [] in let rec proof_extractor vl = function | {ref=Some(Prim _,_)} as pf -> prim_extractor proof_extractor vl pf @@ -255,42 +253,31 @@ let extract_open_proof sign pf = | {ref=Some(Local_constraints lc,[pf])} -> (proof_extractor vl) pf | {ref=None;goal=goal} -> - let n_rels = number_of_rels vl in let visible_rels = map_succeed (fun id -> - match lookup_id id vl with - | GLOBNAME _ -> failwith "caught" - | RELNAME(n,_) -> (n,id)) - (ids_of_sign (evar_hyps goal)) in - let sorted_rels = - Sort.list (fun (n1,_) (n2,_) -> n1>n2) visible_rels in + try let n = list_index id vl in (n,id) + with Not_found -> failwith "caught") + (ids_of_var_context (evar_hyps goal)) in + let sorted_rels = + Sort.list (fun (n1,_) (n2,_) -> n1 > n2 ) visible_rels in let abs_concl = List.fold_right (fun (_,id) concl -> - let (_,ty) = lookup_sign id (evar_hyps goal) in - mkNamedProd id (incast_type ty) concl) + let (c,ty) = lookup_id id (evar_hyps goal) in + mkNamedProd_or_LetIn (id,c,ty) concl) sorted_rels goal.evar_concl in + let ass = Retyping.get_assumption_of goal.evar_env sigma abs_concl in let mv = new_meta() in - open_obligations := (mv,abs_concl):: !open_obligations; + open_obligations := (mv,ass):: !open_obligations; applist(DOP0(Meta mv),List.map (fun (n,_) -> Rel n) sorted_rels) | _ -> anomaly "Bug : a case has been forgotten in proof_extractor" in - let pfterm = proof_extractor (gLOB sign) pf in + let pfterm = proof_extractor [] pf in (pfterm, List.rev !open_obligations) -(* extracts a constr from a proof, and raises an error if the proof is - incomplete *) - -let extract_proof sign pf = - match extract_open_proof sign pf with - | t,[] -> t - | _ -> errorlabstrm "extract_proof" - [< 'sTR "Attempt to save an incomplete proof" >] - - (*********************) (* Tacticals *) (*********************) @@ -554,7 +541,11 @@ let tactic_list_tactic tac gls = -(* solve_subgoal n tac pf_sigma applies the tactic tac at the nth subgoal of +(* solve_subgoal : + (global_constraints ref * goal list * validation -> + global_constraints ref * goal list * validation) -> + (proof_tree sigma -> proof_tree sigma) + solve_subgoal tac pf_sigma applies the tactic tac at the nth subgoal of pf_sigma *) let solve_subgoal tacl pf_sigma = let (sigr,pf) = unpackage pf_sigma in @@ -587,7 +578,7 @@ type pftreestate = { let proof_of_pftreestate pts = pts.tpf let is_top_pftreestate pts = pts.tstack = [] let cursor_of_pftreestate pts = List.map fst pts.tstack -let evc_of_pftreestate pts = pts.tpfsigma +let evc_of_pftreestate pts = ts_it pts.tpfsigma let top_goal_of_pftreestate pts = { it = goal_of_proof pts.tpf; sigma = pts.tpfsigma } @@ -649,12 +640,12 @@ let change_constraints_pftreestate newgc pts = (* solve the nth subgoal with tactic tac *) let solve_nth_pftreestate n tac pts = - let pf = pts.tpf in - let rslts = - solve_subgoal (theni_tac n tac) {it = pts.tpf;sigma = pts.tpfsigma} in - { tpf = rslts.it; - tpfsigma = rslts.sigma; - tstack = pts.tstack } + let rslts = + solve_subgoal (theni_tac n tac) {it = pts.tpf;sigma = pts.tpfsigma} + in + { tpf = rslts.it; + tpfsigma = rslts.sigma; + tstack = pts.tstack } let solve_pftreestate = solve_nth_pftreestate 1 @@ -678,15 +669,19 @@ let mk_pftreestate g = proof is not complete or the state does not correspond to the head of the proof-tree *) +let extract_open_pftreestate pts = + extract_open_proof (ts_it pts.tpfsigma) pts.tpf + let extract_pftreestate pts = if pts.tstack <> [] then errorlabstrm "extract_pftreestate" [< 'sTR"Cannot extract from a proof-tree in which we have descended;" ; 'sPC; 'sTR"Please ascend to the root" >]; - let env = pts.tpf.goal.evar_env in - let hyps = Environ.var_context env in - strong whd_betaiotaevar env (ts_it pts.tpfsigma) - (extract_proof hyps pts.tpf) + let pfterm,subgoals = extract_open_pftreestate pts in + if subgoals <> [] then + errorlabstrm "extract_proof" + [< 'sTR "Attempt to save an incomplete proof" >]; + strong whd_betaiotaevar pts.tpf.goal.evar_env (ts_it pts.tpfsigma) pfterm (* Focus on the first leaf proof in a proof-tree state *) @@ -811,18 +806,22 @@ let pr_rule = function | Context ctxt -> pr_ctxt ctxt | Local_constraints lc -> [< 'sTR"Local constraint change" >] -let thin_sign osign sign = - sign_it - (fun id ty nsign -> - if (not (mem_sign osign id)) - or (id,ty) <> (lookup_sign id osign) (* Hum, egalité sur les types *) - then add_sign (id,ty) nsign - else nsign) sign nil_sign +exception Different + +(* We remove from the var context of env what is already in osign *) +let thin_sign osign env = + process_var_context + (fun env (id,c,ty as d) -> + try + if lookup_id id osign = (c,ty) then env + else raise Different + with Not_found | Different -> push_var d env) + env let rec print_proof sigma osign pf = let {evar_env=env; evar_concl=cl; evar_info=info; evar_body=body} = pf.goal in - let env' = change_hyps (fun sign -> thin_sign osign sign) env in + let env' = thin_sign osign env in match pf.ref with | None -> hOV 0 [< pr_seq {evar_env=env'; evar_concl=cl; diff --git a/proofs/refiner.mli b/proofs/refiner.mli index 5c2b24d7f..8546b0c73 100644 --- a/proofs/refiner.mli +++ b/proofs/refiner.mli @@ -39,10 +39,6 @@ val frontier : transformation_tactic val list_pf : proof_tree -> goal list val unTAC : tactic -> goal sigma -> proof_tree sigma -val extract_open_proof : - typed_type signature -> proof_tree -> constr * (int * constr) list - - (*s Tacticals. *) (* [tclIDTAC] is the identity tactic *) @@ -114,7 +110,7 @@ type pftreestate val proof_of_pftreestate : pftreestate -> proof_tree val cursor_of_pftreestate : pftreestate -> int list val is_top_pftreestate : pftreestate -> bool -val evc_of_pftreestate : pftreestate -> global_constraints +val evc_of_pftreestate : pftreestate -> evar_declarations val top_goal_of_pftreestate : pftreestate -> goal sigma val nth_goal_of_pftreestate : int -> pftreestate -> goal sigma @@ -127,6 +123,7 @@ val solve_pftreestate : tactic -> pftreestate -> pftreestate val weak_undo_pftreestate : pftreestate -> pftreestate val mk_pftreestate : goal -> pftreestate +val extract_open_pftreestate : pftreestate -> constr * (int * typed_type) list val extract_pftreestate : pftreestate -> constr val first_unproven : pftreestate -> pftreestate val last_unproven : pftreestate -> pftreestate diff --git a/proofs/tacinterp.ml b/proofs/tacinterp.ml index c95424fb1..cb8eb0c45 100644 --- a/proofs/tacinterp.ml +++ b/proofs/tacinterp.ml @@ -45,12 +45,6 @@ let constr_of_Constr = function | Constr c -> c | _ -> anomalylabstrm "constr_of_Constr" [<'sTR "Not a CONSTR tactic_arg">] -(* Transforms a type_judgment signature into a (string * constr) list *) -let make_hyps hyps = - let lid = List.map string_of_id (ids_of_sign hyps) - and lhyp = List.map body_of_type (vals_of_sign hyps) in - List.rev (List.combine lid lhyp) - (* Extracted the constr list from lfun *) let rec constr_list goalopt = function | (str,VArg(Constr c))::tl -> (id_of_string str,c)::(constr_list goalopt tl) @@ -62,7 +56,7 @@ let rec constr_list goalopt = function (match goalopt with | None -> constr_list goalopt tl | Some goal -> - if List.mem_assoc (string_of_id id) (make_hyps (pf_hyps goal)) then + if mem_var_context id (pf_hyps goal) then (id_of_string str,VAR id)::(constr_list goalopt tl) else constr_list goalopt tl)) @@ -287,12 +281,11 @@ let apply_matching pat csr = PatternMatchingFailure -> raise No_match (* Tries to match one hypothesis with a list of hypothesis patterns *) -let apply_one_hyp_context lmatch mhyps (idhyp,hyp) = +let apply_one_hyp_context lmatch mhyps (idhyp,bodyopt,hyp) = let rec apply_one_hyp_context_rec (idhyp,hyp) mhyps_acc = function | (Hyp (idpat,pat))::tl -> (try - ([idpat,VArg (Identifier - (id_of_string idhyp))],verify_metas_coherence lmatch + ([idpat,VArg (Identifier idhyp)],verify_metas_coherence lmatch (Pattern.matches pat hyp),mhyps_acc@tl) with PatternMatchingFailure | Not_coherent_metas -> @@ -307,7 +300,7 @@ let apply_one_hyp_context lmatch mhyps (idhyp,hyp) = apply_one_hyp_context_rec (idhyp,hyp) (mhyps_acc@[NoHypId pat]) tl) | [] -> raise No_match in - apply_one_hyp_context_rec (idhyp,hyp) [] mhyps + apply_one_hyp_context_rec (idhyp,body_of_type hyp) [] mhyps (* Prepares the lfun list for constr substitutions *) let rec make_subs_list = function @@ -473,7 +466,7 @@ and match_context_interp evc env lfun lmatch goalopt ast lmr = (try eval_with_fail (val_interp (evc,env,lfun,lmatch,Some goal)) t goal with No_match -> apply_match_context evc env lfun lmatch goal tl) | (Pat (mhyps,mgoal,mt))::tl -> - let hyps = make_hyps (pf_hyps goal) + let hyps = pf_hyps goal and concl = pf_concl goal in (try (let lgoal = apply_matching mgoal concl in diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 9471b5d25..1488b06de 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -40,19 +40,27 @@ let pf_env gls = (sig_it gls).evar_env let pf_hyps gls = var_context (sig_it gls).evar_env let pf_concl gls = (sig_it gls).evar_concl - +(* let pf_untyped_hyps gls = let sign = Environ.var_context (pf_env gls) in map_sign_typ (fun x -> body_of_type x) sign +*) +let pf_hyps_types gls = + let sign = Environ.var_context (pf_env gls) in + List.map (fun (id,_,x) -> (id,body_of_type x)) sign + +let pf_nth_hyp_id gls n = let (id,c,t) = List.nth (pf_hyps gls) (n-1) in id -let pf_nth_hyp gls n = nth_sign (pf_untyped_hyps gls) n +let pf_last_hyp gl = List.hd (pf_hyps gl) -let pf_get_hyp gls id = +let pf_get_hyp_typ gls id = try - snd (lookup_sign id (pf_untyped_hyps gls)) + body_of_type (snd (lookup_id id (pf_hyps gls))) with Not_found -> error ("No such hypothesis : " ^ (string_of_id id)) +let pf_ids_of_hyps gls = ids_of_var_context (var_context (pf_env gls)) + let pf_ctxt gls = get_ctxt (sig_it gls) let pf_interp_constr gls c = @@ -135,6 +143,7 @@ let solve_pftreestate = solve_pftreestate let weak_undo_pftreestate = weak_undo_pftreestate let mk_pftreestate = mk_pftreestate let extract_pftreestate = extract_pftreestate +let extract_open_pftreestate = extract_open_pftreestate let first_unproven = first_unproven let last_unproven = last_unproven let nth_unproven = nth_unproven @@ -218,6 +227,12 @@ let unTAC = unTAC let refiner = refiner +(* This does not check that the variable name is not here *) +let unsafe_introduction id = + without_check + (refiner (Prim { name = Intro; newids = [id]; + hypspecs = []; terms = []; params = [] })) + let introduction id pf = refiner (Prim { name = Intro; newids = [id]; hypspecs = []; terms = []; params = [] }) pf @@ -259,7 +274,7 @@ let mutual_cofix lf lar pf = let rename_bound_var_goal gls = let { evar_env = env; evar_concl = cl } as gl = sig_it gls in - let ids = ids_of_sign (Environ.var_context env) in + let ids = ids_of_var_context (Environ.var_context env) in convert_concl (rename_bound_var ids cl) gls @@ -473,7 +488,7 @@ open Printer let pr_com sigma goal com = prterm (rename_bound_var - (ids_of_sign (var_context goal.evar_env)) + (ids_of_var_context (var_context goal.evar_env)) (Astterm.interp_constr sigma goal.evar_env com)) let pr_one_binding sigma goal = function diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 60be33174..6ae4f9ae2 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -34,8 +34,11 @@ val apply_sig_tac : val pf_concl : goal sigma -> constr val pf_env : goal sigma -> env val pf_hyps : goal sigma -> var_context -val pf_untyped_hyps : goal sigma -> constr signature -val pf_nth_hyp : goal sigma -> int -> identifier * constr +(*val pf_untyped_hyps : goal sigma -> (identifier * constr) list*) +val pf_hyps_types : goal sigma -> (identifier * constr) list +val pf_nth_hyp_id : goal sigma -> int -> identifier +val pf_last_hyp : goal sigma -> var_declaration +val pf_ids_of_hyps : goal sigma -> identifier list val pf_ctxt : goal sigma -> ctxtty val pf_global : goal sigma -> identifier -> constr val pf_parse_const : goal sigma -> string -> constr @@ -45,9 +48,9 @@ val pf_execute : goal sigma -> constr -> unsafe_judgment val hnf_type_of : goal sigma -> constr -> constr val pf_interp_constr : goal sigma -> Coqast.t -> constr -val pf_interp_type : goal sigma -> Coqast.t -> constr +val pf_interp_type : goal sigma -> Coqast.t -> constr -val pf_get_hyp : goal sigma -> identifier -> constr +val pf_get_hyp_typ : goal sigma -> identifier -> constr val pf_reduction_of_redexp : goal sigma -> red_expr -> constr -> constr @@ -88,7 +91,7 @@ type pftreestate val proof_of_pftreestate : pftreestate -> proof_tree val cursor_of_pftreestate : pftreestate -> int list val is_top_pftreestate : pftreestate -> bool -val evc_of_pftreestate : pftreestate -> global_constraints +val evc_of_pftreestate : pftreestate -> evar_declarations val top_goal_of_pftreestate : pftreestate -> goal sigma val nth_goal_of_pftreestate : int -> pftreestate -> goal sigma val traverse : int -> pftreestate -> pftreestate @@ -96,6 +99,7 @@ val weak_undo_pftreestate : pftreestate -> pftreestate val solve_nth_pftreestate : int -> tactic -> pftreestate -> pftreestate val solve_pftreestate : tactic -> pftreestate -> pftreestate val mk_pftreestate : goal -> pftreestate +val extract_open_pftreestate : pftreestate -> constr * (int * typed_type) list val extract_pftreestate : pftreestate -> constr val first_unproven : pftreestate -> pftreestate val last_unproven : pftreestate -> pftreestate diff --git a/tactics/auto.ml b/tactics/auto.ml index 7de32f90f..9b0d24632 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -239,9 +239,10 @@ let make_resolves name eap (c,cty) = ents (* used to add an hypothesis to the local hint database *) -let make_resolve_hyp hname htyp = +let make_resolve_hyp (hname,_,htyp) = try - [make_apply_entry (true, Options.is_verbose()) hname (mkVar hname, htyp)] + [make_apply_entry (true, Options.is_verbose()) hname + (mkVar hname, body_of_type htyp)] with | Failure _ -> [] | e when Logic.catchable_exception e -> anomaly "make_resolve_hyp" @@ -596,12 +597,10 @@ let unify_resolve (c,clenv) gls = h_simplest_apply c gls (* builds a hint database from a constr signature *) -(* typically used with (lid, ltyp) = pf_untyped_hyps <some goal> *) +(* typically used with (lid, ltyp) = pf_hyps_types <some goal> *) let make_local_hint_db sign = - let hintlist = - list_map_append2 make_resolve_hyp - (ids_of_sign sign) (vals_of_sign sign) in + let hintlist = list_map_append make_resolve_hyp sign in Hint_db.add_list hintlist Hint_db.empty @@ -617,8 +616,7 @@ let rec trivial_fail_db db_list local_db gl = let intro_tac = tclTHEN intro (fun g'-> - let (hn, htyp) = hd_sign (pf_untyped_hyps g') in - let hintl = make_resolve_hyp hn htyp in + let hintl = make_resolve_hyp (pf_last_hyp g') in trivial_fail_db db_list (Hint_db.add_list hintl local_db) g') in tclFIRST @@ -668,14 +666,13 @@ let trivial dbnames gl = error ("Trivial: "^x^": No such Hint database")) ("core"::dbnames) in - tclTRY - (trivial_fail_db db_list (make_local_hint_db (pf_untyped_hyps gl))) gl + tclTRY (trivial_fail_db db_list (make_local_hint_db (pf_hyps gl))) gl let full_trivial gl = let dbnames = stringmap_dom !searchtable in let dbnames = list_subtract dbnames ["v62"] in let db_list = List.map (fun x -> searchtable_map x) dbnames in - tclTRY (trivial_fail_db db_list (make_local_hint_db (pf_untyped_hyps gl))) gl + tclTRY (trivial_fail_db db_list (make_local_hint_db (pf_hyps gl))) gl let dyn_trivial = function | [] -> trivial [] @@ -733,28 +730,26 @@ let rec search_gen decomp n db_list local_db extra_sign goal = (fun id -> tclTHEN (decomp_unary_term (VAR id)) (tclTHEN (clear_one id) - (search_gen decomp p db_list local_db nil_sign))) - (ids_of_sign (pf_hyps goal))) + (search_gen decomp p db_list local_db []))) + (pf_ids_of_hyps goal)) in let intro_tac = tclTHEN intro (fun g' -> - let (hid,htyp) = hd_sign (pf_untyped_hyps g') in + let (hid,_,htyp as d) = pf_last_hyp g' in let hintl = try [make_apply_entry (true,Options.is_verbose()) - hid (mkVar hid,htyp)] + hid (mkVar hid,body_of_type htyp)] with Failure _ -> [] in - search_gen decomp n db_list - (Hint_db.add_list hintl local_db) - (add_sign (hid,htyp) nil_sign) g') + search_gen decomp n db_list (Hint_db.add_list hintl local_db) [d] g') in let rec_tacs = List.map (fun ntac -> tclTHEN ntac - (search_gen decomp (n-1) db_list local_db nil_sign)) + (search_gen decomp (n-1) db_list local_db empty_var_context)) (possible_resolve db_list local_db (pf_concl goal)) in tclFIRST (assumption::(decomp_tacs@(intro_tac::rec_tacs))) goal @@ -774,7 +769,7 @@ let auto n dbnames gl = error ("Auto: "^x^": No such Hint database")) ("core"::dbnames) in - let hyps = (pf_untyped_hyps gl) in + let hyps = pf_hyps gl in tclTRY (search n db_list (make_local_hint_db hyps) hyps) gl let default_auto = auto !default_search_depth [] @@ -783,7 +778,7 @@ let full_auto n gl = let dbnames = stringmap_dom !searchtable in let dbnames = list_subtract dbnames ["v62"] in let db_list = List.map (fun x -> searchtable_map x) dbnames in - let hyps = (pf_untyped_hyps gl) in + let hyps = pf_hyps gl in tclTRY (search n db_list (make_local_hint_db hyps) hyps) gl let default_full_auto gl = full_auto !default_search_depth gl @@ -818,7 +813,7 @@ let h_auto = hide_tactic "Auto" dyn_auto let default_search_decomp = ref 1 let destruct_auto des_opt n gl = - let hyps = pf_untyped_hyps gl in + let hyps = pf_hyps gl in search_gen des_opt n [searchtable_map "core"] (make_local_hint_db hyps) hyps gl @@ -850,23 +845,23 @@ type autoArguments = let keepAfter tac1 tac2 = (tclTHEN tac1 - (function g -> tac2 (make_sign [hd_sign (pf_untyped_hyps g)]) g)) + (function g -> tac2 [pf_last_hyp g] g)) let compileAutoArg contac = function | Destructing -> (function g -> - let ctx =(pf_untyped_hyps g) in + let ctx = pf_hyps g in tclFIRST - (List.map2 - (fun id typ -> - if (Hipattern.is_conjunction (hd_of_prod typ)) then + (List.map + (fun (id,_,typ) -> + if (Hipattern.is_conjunction (hd_of_prod (body_of_type typ))) + then (tclTHEN (tclTHEN (simplest_elim (mkVar id)) (clear_one id)) contac) else - tclFAIL 0) - (ids_of_sign ctx) (vals_of_sign ctx)) g) + tclFAIL 0) ctx) g) | UsingTDB -> (tclTHEN (Tacticals.tryAllClauses @@ -884,8 +879,9 @@ let rec super_search n db_list local_db argl goal = :: (tclTHEN intro (fun g -> - let (hid,htyp) = hd_sign (pf_untyped_hyps g) in - let hintl = make_resolves hid (true,false) (mkVar hid,htyp) in + let (hid,_,htyp) = pf_last_hyp g in + let hintl = + make_resolves hid (true,false) (mkVar hid,body_of_type htyp) in super_search n db_list (Hint_db.add_list hintl local_db) argl g)) :: @@ -901,9 +897,11 @@ let rec super_search n db_list local_db argl goal = let search_superauto n ids argl g = let sigma = List.fold_right - (fun id -> add_sign (id,pf_type_of g (pf_global g id))) - ids nil_sign in - let hyps = concat_sign (pf_untyped_hyps g) sigma in + (fun id -> add_var_decl + (id,Retyping.get_assumption_of (pf_env g) (project g) + (pf_type_of g (pf_global g id)))) + ids empty_var_context in + let hyps = List.append (pf_hyps g) sigma in super_search n [Stringmap.find "core" !searchtable] (make_local_hint_db hyps) argl g diff --git a/tactics/auto.mli b/tactics/auto.mli index 50460de17..b65d2ea21 100644 --- a/tactics/auto.mli +++ b/tactics/auto.mli @@ -85,8 +85,7 @@ val make_resolves : Never raises an User_exception; If the hyp cannot be used as a Hint, the empty list is returned. *) -val make_resolve_hyp : identifier -> constr - -> (constr_label * pri_auto_tactic) list +val make_resolve_hyp : var_declaration -> (constr_label * pri_auto_tactic) list (* [make_extern name pri pattern tactic_ast] *) @@ -97,7 +96,7 @@ val make_extern : (* Create a Hint database from the pairs (name, constr). Useful to take the current goal hypotheses as hints *) -val make_local_hint_db : constr signature -> Hint_db.t +val make_local_hint_db : var_context -> Hint_db.t val priority : (int * 'a) list -> 'a list diff --git a/tactics/eauto.ml b/tactics/eauto.ml index 74d49114b..25bb62ab4 100644 --- a/tactics/eauto.ml +++ b/tactics/eauto.ml @@ -21,13 +21,13 @@ let e_give_exact c gl = tclTHEN (unify (pf_type_of gl c)) (Tactics.exact c) gl let assumption id = e_give_exact (VAR id) let e_assumption gl = - tclFIRST (List.map assumption (ids_of_sign (pf_untyped_hyps gl))) gl + tclFIRST (List.map assumption (pf_ids_of_hyps gl)) gl let e_give_exact_constr = hide_constr_tactic "EExact" e_give_exact let registered_e_assumption gl = tclFIRST (List.map (fun id gl -> e_give_exact_constr (VAR id) gl) - (ids_of_sign (pf_untyped_hyps gl))) gl + (pf_ids_of_hyps gl)) gl let e_resolve_with_bindings_tac (c,lbind) gl = let (wc,kONT) = startWalk gl in @@ -54,9 +54,9 @@ let vernac_e_resolve_constr = let one_step l gl = [Tactics.intro] @ (List.map e_resolve_constr (List.map (fun x -> VAR x) - (ids_of_sign (pf_untyped_hyps gl)))) + (pf_ids_of_hyps gl))) @ (List.map e_resolve_constr l) - @ (List.map assumption (ids_of_sign (pf_untyped_hyps gl))) + @ (List.map assumption (pf_ids_of_hyps gl)) let rec prolog l n gl = if n <= 0 then error "prolog - failure"; @@ -152,8 +152,8 @@ let rec e_trivial_fail_db db_list local_db goal = registered_e_assumption :: (tclTHEN Tactics.intro (function g'-> - let (hn,htyp) = hd_sign (pf_untyped_hyps g') in - let hintl = make_resolve_hyp hn htyp in + let d = pf_last_hyp g' in + let hintl = make_resolve_hyp d in (e_trivial_fail_db db_list (Hint_db.add_list hintl local_db) g'))) :: (e_trivial_resolve db_list local_db (pf_concl goal)) @@ -223,14 +223,13 @@ let rec e_search n db_list local_db lg = (assumption_tac_list id) (e_search n db_list local_db) gls) - (ids_of_sign (pf_untyped_hyps g)) + (pf_ids_of_hyps g) in let intro_tac = apply_tac_list (tclTHEN Tactics.intro (fun g' -> - let (hid,htyp) = hd_sign (pf_untyped_hyps g') in - let hintl = make_resolve_hyp hid htyp in + let hintl = make_resolve_hyp (pf_last_hyp g') in (tactic_list_tactic (e_search n db_list (Hint_db.add_list hintl local_db))) g')) @@ -249,7 +248,7 @@ let rec e_search n db_list local_db lg = let e_search_auto n db_list gl = tactic_list_tactic - (e_search n db_list (make_local_hint_db (pf_untyped_hyps gl))) + (e_search n db_list (make_local_hint_db (pf_hyps gl))) gl let eauto n dbnames = @@ -268,7 +267,7 @@ let full_eauto n gl = let dbnames = stringmap_dom !searchtable in let dbnames = list_subtract dbnames ["v62"] in let db_list = List.map (fun x -> Stringmap.find x !searchtable) dbnames in - let local_db = make_local_hint_db (pf_untyped_hyps gl) in + let local_db = make_local_hint_db (pf_hyps gl) in tclTRY (tclCOMPLETE (e_search_auto n db_list)) gl let default_full_auto gl = full_auto !default_search_depth gl diff --git a/tactics/equality.ml b/tactics/equality.ml index 22bc94b69..ae5bed674 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -47,10 +47,12 @@ let general_rewrite_bindings lft2rgt (c,l) gl = let hdcncls = string_head hdcncl in let elim = if lft2rgt then - pf_global gl (id_of_string - (hdcncls^(suff gl (pf_concl gl))^"_r")) + pf_global gl + (id_of_string + (hdcncls^(Declare.elimination_suffix (sort_of_goal gl))^"_r")) else - pf_global gl (id_of_string (hdcncls^(suff gl (pf_concl gl)))) + pf_global gl + (id_of_string (hdcncls^(Declare.elimination_suffix (sort_of_goal gl)))) in tclNOTSAMEGOAL (general_elim (c,l) (elim,[])) gl (* was tclWEAK_PROGRESS which only fails for tactics generating one subgoal @@ -445,8 +447,8 @@ let discriminable env sigma t1 t2 = the continuation then constructs the case-split. *) -let push_rel_type sigma (na,t) env = - push_rel (na,make_typed t (get_sort_of env sigma t)) env +let push_rel_type sigma (na,c,t) env = + push_rel (na,c,t) env let push_rels vars env = List.fold_right (fun nvar env -> push_rel_type Evd.empty nvar env) vars env @@ -466,7 +468,7 @@ let descend_then sigma env head dirn = let p = lam_it (lift (mis_nrealargs mispec) resty) aritysign in let build_branch i = let result = if i = dirn then dirnval else dfltval in - it_lambda_name env result cstr.(i-1).cs_args + it_mkLambda_or_LetIn_name env result cstr.(i-1).cs_args in mkMutCase (make_default_case_info mispec) p head (List.map build_branch (interval 1 (mis_nconstr mispec))))) @@ -510,7 +512,7 @@ let construct_discriminator sigma env dirn c sort = let cstrs = get_constructors indf in let build_branch i = let endpt = if i = dirn then true_0 else false_0 in - lam_it endpt cstrs.(i-1).cs_args + it_mkLambda_or_LetIn endpt cstrs.(i-1).cs_args in let build_match = mkMutCase (make_default_case_info mispec) p c @@ -608,7 +610,7 @@ let discr id gls = | Inl(cpath,MutConstruct(_,dirn),_) -> let e = pf_get_new_id (id_of_string "ee") gls in let e_env = - push_var (e,assumption_of_judgment env sigma tj) env + push_var_decl (e,assumption_of_judgment env sigma tj) env in let discriminator = build_discriminator sigma e_env dirn (VAR e) sort cpath in @@ -714,7 +716,7 @@ let make_tuple env sigma (rterm,rty) lind = let a = type_of env sigma (Rel lind) in (* We replace (Rel lind) by (Rel 1) in rty then abstract on (na:a) *) let rty' = substnl [Rel 1] lind rty in - let na = fst (lookup_rel lind env) in + let na = fst (lookup_rel_type lind env) in let p = mkLambda na a rty' in (applist(exist_term,[a;p;(Rel lind);rterm]), applist(sig_term,[a;p])) @@ -888,7 +890,7 @@ let inj id gls = | Inr posns -> let e = pf_get_new_id (id_of_string "e") gls in let e_env = - push_var (e,assumption_of_judgment env sigma tj) env + push_var_decl (e,assumption_of_judgment env sigma tj) env in let injectors = map_succeed @@ -954,7 +956,7 @@ let decompEqThen ntac id gls = | Inl(cpath,MutConstruct(_,dirn),_) -> let e = pf_get_new_id (id_of_string "e") gls in let e_env = - push_var (e,assumption_of_judgment env sigma tj) env in + push_var_decl (e,assumption_of_judgment env sigma tj) env in let discriminator = build_discriminator sigma e_env dirn (VAR e) sort cpath in let (pf, absurd_term) = @@ -966,7 +968,7 @@ let decompEqThen ntac id gls = | Inr posns -> (let e = pf_get_new_id (id_of_string "e") gls in let e_env = - push_var (e,assumption_of_judgment env sigma tj) env in + push_var_decl (e,assumption_of_judgment env sigma tj) env in let injectors = map_succeed (fun (cpath,t1_0,t2_0) -> diff --git a/tactics/inv.ml b/tactics/inv.ml index 385cd574c..39254ace0 100644 --- a/tactics/inv.ml +++ b/tactics/inv.ml @@ -66,16 +66,9 @@ let dest_match_eq gls eqn = let implicit = Sort implicit_sort -let change_sign env (vars,rels) = - let env' = change_hyps (fun _ -> vars) env in - List.fold_left - (fun env (n,t) -> - let tt = execute_type env Evd.empty t in push_rel (n,tt) env) - env' rels - (* Environment management *) let push_rel_type sigma (na,t) env = - push_rel (na,make_typed t (get_sort_of env sigma t)) env + push_rel_decl (na,make_typed t (get_sort_of env sigma t)) env let push_rels vars env = List.fold_left (fun env nvar -> push_rel_type Evd.empty nvar env) env vars @@ -111,7 +104,7 @@ let make_inv_predicate env sigma ind id status concl = let p = make_arity env true indf sort in abstract_list_all env sigma p concl (realargs@[VAR id]) in - let hyps,_ = decompose_lam pred in + let hyps,_ = decompose_lam_n (nrealargs+1) pred in let c3 = whd_beta (applist (pred,rel_list nrealargs (nrealargs +1))) in (hyps,c3) @@ -210,7 +203,7 @@ let generalizeRewriteIntros tac depids id gls = let var_occurs_in_pf gl id = occur_var id (pf_concl gl) or - exists_sign (fun _ t -> occur_var id t) (pf_untyped_hyps gl) + List.exists (fun (_,t) -> occur_var id t) (pf_hyps_types gl) let split_dep_and_nodep idl gl = (List.filter (var_occurs_in_pf gl) idl, diff --git a/tactics/leminv.ml b/tactics/leminv.ml index 60a4cd364..9e4386f1d 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -28,9 +28,9 @@ open Inv let not_work_message = "tactic fails to build the inversion lemma, may be because the predicate has arguments that depend on other arguments" -let no_inductive_inconstr ass constr = +let no_inductive_inconstr env constr = [< 'sTR "Cannot recognize an inductive predicate in "; - prterm_env (Environ.context ass) constr; + prterm_env env constr; 'sTR "."; 'sPC; 'sTR "If there is one, may be the structure of the arity"; 'sPC; 'sTR "or of the type of constructors"; 'sPC; 'sTR "is hidden by constant definitions." >] @@ -80,16 +80,18 @@ let no_inductive_inconstr ass constr = let thin_ids (hyps,vars) = fst - (it_sign - (fun ((ids,globs) as sofar) id a -> + (List.fold_left + (fun ((ids,globs) as sofar) (id,c,a) -> if List.mem id globs then - (id::ids,(global_vars a)@globs) + match c with + | None -> (id::ids,(global_vars a)@globs) + | Some body -> (id::ids,(global_vars body)@(global_vars a)@globs) else sofar) ([],vars) hyps) (* returns the sub_signature of sign corresponding to those identifiers that * are not global. *) - +(* let get_local_sign sign = let lid = ids_of_sign sign in let globsign = Global.var_context() in @@ -100,13 +102,13 @@ let get_local_sign sign = res_sign in List.fold_right add_local lid nil_sign - +*) (* returs the identifier of lid that was the latest declared in sign. * (i.e. is the identifier id of lid such that * sign_length (sign_prefix id sign) > sign_length (sign_prefix id' sign) > * for any id'<>id in lid). * it returns both the pair (id,(sign_prefix id sign)) *) - +(* let max_prefix_sign lid sign = let rec max_rec (resid,prefix) = function | [] -> (resid,prefix) @@ -120,14 +122,14 @@ let max_prefix_sign lid sign = match lid with | [] -> nil_sign | id::l -> snd (max_rec (id, sign_prefix id sign) l) - +*) let rec add_prods_sign env sigma t = match kind_of_term (whd_betadeltaiota env sigma t) with | IsProd (na,c1,b) -> let id = Environ.id_of_name_using_hdchar env t na in let b'= subst1 (VAR id) b in - let j = make_typed c1 (Retyping.get_sort_of env sigma c1) in - add_prods_sign (Environ.push_var (id,j) env) sigma b' + let j = Retyping.get_assumption_of env sigma c1 in + add_prods_sign (Environ.push_var_decl (id,j) env) sigma b' | _ -> (env,t) (* [dep_option] indicates wether the inversion lemma is dependent or not. @@ -145,7 +147,7 @@ let rec add_prods_sign env sigma t = let compute_first_inversion_scheme env sigma ind sort dep_option = let indf,realargs = dest_ind_type ind in - let allvars = ids_of_sign (var_context env) in + let allvars = ids_of_context env in let p = next_ident_away (id_of_string "P") allvars in let pty,goal = if dep_option then @@ -158,20 +160,18 @@ let compute_first_inversion_scheme env sigma ind sort dep_option = let i = mkAppliedInd ind in let ivars = global_vars i in let revargs,ownsign = - sign_it - (fun id a (revargs,hyps) -> - if List.mem id ivars then - ((VAR id)::revargs,(Name id,body_of_type a)::hyps) + fold_var_context + (fun env (id,_,_ as d) (revargs,hyps) -> + if List.mem id ivars then ((VAR id)::revargs,add_var d hyps) else (revargs,hyps)) - (var_context env) ([],[]) - in - let pty = it_prod_name env (mkSort sort) ownsign in + env ([],[]) in + let pty = it_mkNamedProd_or_LetIn (mkSort sort) ownsign in let goal = mkArrow i (applist(VAR p, List.rev revargs)) in (pty,goal) in let npty = nf_betadeltaiota env sigma pty in let nptyj = make_typed npty (Retyping.get_sort_of env sigma npty) in - let extenv = push_var (p,nptyj) env in + let extenv = push_var_decl (p,nptyj) env in extenv, goal (* [inversion_scheme sign I] @@ -189,7 +189,7 @@ let inversion_scheme env sigma t sort dep_option inv_op = errorlabstrm "inversion_scheme" (no_inductive_inconstr env i) in let (invEnv,invGoal) = compute_first_inversion_scheme env sigma ind sort dep_option in - assert (list_subset (global_vars invGoal)(ids_of_sign (var_context invEnv))); + assert (list_subset (global_vars invGoal) (ids_of_var_context (var_context invEnv))); (* errorlabstrm "lemma_inversion" [< 'sTR"Computed inversion goal was not closed in initial signature" >]; @@ -198,24 +198,23 @@ let inversion_scheme env sigma t sort dep_option inv_op = let pfs = solve_pftreestate (tclTHEN intro (onLastHyp (compose inv_op out_some))) pfs in - let pf = proof_of_pftreestate pfs in - let (pfterm,meta_types) = - Refiner.extract_open_proof (var_context pf.goal.evar_env) pf in + let (pfterm,meta_types) = extract_open_pftreestate pfs in + let global_var_context = Global.var_context () in let ownSign = - sign_it - (fun id ty sign -> - if mem_sign (Global.var_context()) id then sign - else ((Name id,body_of_type ty)::sign)) - (var_context invEnv) [] in + fold_var_context + (fun env (id,_,_ as d) sign -> + if mem_var_context id global_var_context then sign + else add_var d sign) + invEnv empty_var_context in let (_,ownSign,mvb) = List.fold_left (fun (avoid,sign,mvb) (mv,mvty) -> let h = next_ident_away (id_of_string "H") avoid in - (h::avoid, (Name h,mvty)::sign, (mv,VAR h)::mvb)) - (ids_of_sign (var_context invEnv), ownSign, []) + (h::avoid, add_var_decl (h,mvty) sign, (mv,VAR h)::mvb)) + (ids_of_context invEnv, ownSign, []) meta_types in let invProof = - it_lambda_name env (local_strong (whd_meta mvb) pfterm) ownSign in + it_mkNamedLambda_or_LetIn (local_strong (whd_meta mvb) pfterm) ownSign in invProof let add_inversion_lemma name env sigma t sort dep inv_op = @@ -232,8 +231,7 @@ let add_inversion_lemma name env sigma t sort dep inv_op = let inversion_lemma_from_goal n na id sort dep_option inv_op = let pts = get_pftreestate() in let gl = nth_goal_of_pftreestate n pts in - let hyps = pf_untyped_hyps gl in - let t = snd (lookup_sign id hyps) in + let t = pf_get_hyp_typ gl id in let env = pf_env gl and sigma = project gl in let fv = global_vars t in (* Pourquoi ??? @@ -329,7 +327,7 @@ let lemInv id c gls = | UserError (a,b) -> errorlabstrm "LemInv" [< 'sTR "Cannot refine current goal with the lemma "; - prterm_env (Global.context()) c >] + prterm_env (Global.env()) c >] let useInversionLemma = let gentac = diff --git a/tactics/refine.ml b/tactics/refine.ml index f953d349e..8a43cff93 100644 --- a/tactics/refine.ml +++ b/tactics/refine.ml @@ -120,7 +120,7 @@ let replace_in_array env a = let fresh env n = let id = match n with Name x -> x | _ -> id_of_string "_" in - next_global_ident_away id (ids_of_sign (var_context env)) + next_global_ident_away id (ids_of_var_context (var_context env)) let rec compute_metamap env = function (* le terme est directement une preuve *) @@ -140,8 +140,8 @@ let rec compute_metamap env = function | DOP2(Lambda,c1,DLAM(name,c2)) as c -> let v = fresh env name in (** let tj = ise_resolve_type true empty_evd [] (gLOB sign) c1 in **) - let tj = execute_type env Evd.empty c1 in - let env' = push_var (v,tj) env in + let tj = Retyping.get_assumption_of env Evd.empty c1 in + let env' = push_var_decl (v,tj) env in begin match compute_metamap env' (subst1 (VAR v) c2) with (* terme de preuve complet *) | TH (_,_,[]) -> TH (c,[],[]) @@ -167,7 +167,7 @@ let rec compute_metamap env = function let vi = List.rev (List.map (fresh env) fi) in let env' = List.fold_left - (fun env (v,ar) -> push_var (v,execute_type env Evd.empty ar) env) + (fun env (v,ar) -> push_var_decl (v,Retyping.get_assumption_of env Evd.empty ar) env) env (List.combine vi (Array.to_list ai)) in diff --git a/tactics/tacticals.ml b/tactics/tacticals.ml index 6d9aca72c..a37a395dd 100644 --- a/tactics/tacticals.ml +++ b/tactics/tacticals.ml @@ -52,7 +52,7 @@ let tclMAP tacfun l = (* apply a tactic to the nth element of the signature *) let tclNTH_HYP m (tac : constr->tactic) gl = - tac (try VAR(fst(nth_sign (pf_untyped_hyps gl) m)) + tac (try VAR(let (id,_,_) = List.nth (pf_hyps gl) (m-1) in id) with Failure _ -> error "No such assumption") gl (* apply a tactic to the last element of the signature *) @@ -65,10 +65,10 @@ let tclTRY_sign (tac : constr->tactic) sign gl = | [s] -> tac (VAR(s)) (* added in order to get useful error messages *) | (s::sl) -> tclORELSE (tac (VAR(s))) (arec sl) in - arec (ids_of_sign sign) gl + arec (ids_of_var_context sign) gl let tclTRY_HYPS (tac : constr->tactic) gl = - tclTRY_sign tac (pf_untyped_hyps gl) gl + tclTRY_sign tac (pf_hyps gl) gl (* OR-branch *) let tryClauses tac = @@ -101,10 +101,10 @@ let nth_clause n gl = if n = 0 then None else if n < 0 then - let id = List.nth (ids_of_sign (pf_untyped_hyps gl)) (-n-1) in + let id = List.nth (pf_ids_of_hyps gl) (-n-1) in Some id else - let id = List.nth (ids_of_sign (pf_untyped_hyps gl)) (n-1) in + let id = List.nth (pf_ids_of_hyps gl) (n-1) in Some id (* Gets the conclusion or the type of a given hypothesis *) @@ -112,7 +112,7 @@ let nth_clause n gl = let clause_type cls gl = match cls with | None -> pf_concl gl - | Some id -> pf_get_hyp gl id + | Some id -> pf_get_hyp_typ gl id (* Functions concerning matching of clausal environments *) @@ -132,7 +132,7 @@ let onCL cfind cltac gl = cltac (cfind gl) gl (* Create a clause list with all the hypotheses from the context *) -let allHyps gl = (List.map in_some (ids_of_sign (pf_untyped_hyps gl))) +let allHyps gl = List.map in_some (pf_ids_of_hyps gl) (* Create a clause list with all the hypotheses from the context, occuring @@ -140,20 +140,19 @@ let allHyps gl = (List.map in_some (ids_of_sign (pf_untyped_hyps gl))) let afterHyp id gl = List.map in_some - (fst (list_splitby (fun hyp -> hyp = id) - (ids_of_sign (pf_untyped_hyps gl)))) + (fst (list_splitby (fun hyp -> hyp = id) (pf_ids_of_hyps gl))) (* Create a singleton clause list with the last hypothesis from then context *) let lastHyp gl = - let (id,_) = hd_sign (pf_untyped_hyps gl) in [(Some id)] + let id = List.hd (pf_ids_of_hyps gl) in [(Some id)] (* Create a clause list with the n last hypothesis from then context *) let nLastHyps n gl = let ids = - try list_firstn n (ids_of_sign (pf_untyped_hyps gl)) + try list_firstn n (pf_ids_of_hyps gl) with Failure "firstn" -> error "Not enough hypotheses in the goal" in List.map in_some ids @@ -162,7 +161,7 @@ let nLastHyps n gl = (* A clause list with the conclusion and all the hypotheses *) let allClauses gl = - let ids = ids_of_sign(pf_untyped_hyps gl) in + let ids = pf_ids_of_hyps gl in (None::(List.map in_some ids)) let onClause t cls gl = t cls gl @@ -193,7 +192,7 @@ let conclPattern concl pat tacast gl = let ast_bindings = List.map (fun (i,c) -> - (i, Termast.ast_of_constr false (assumptions_for_print []) c)) + (i, Termast.ast_of_constr false (pf_env gl) c)) constr_bindings in let tacast' = Coqast.subst_meta ast_bindings tacast in Tacinterp.interp tacast' gl @@ -285,21 +284,6 @@ let sort_of_goal gl = (* Find the right elimination suffix corresponding to the sort of the goal *) (* c should be of type A1->.. An->B with B an inductive definition *) -let suff gl cl = match hnf_type_of gl cl with - | DOP0(Sort(Type(_))) -> "_rect" - | DOP0(Sort(Prop(Null))) -> "_ind" - | DOP0(Sort(Prop(Pos))) -> "_rec" - | _ -> anomaly "goal should be a type" - -(* Look up function for the default elimination constant *) - -let lookup_eliminator sign path suff = - let name = id_of_string ((string_of_id (basename path)) ^ suff) in - try - Declare.global_reference (kind_of_path path) name - with UserError _ -> - VAR(fst(lookup_glob name (gLOB sign))) - let last_arg = function | DOPN(AppL,cl) -> cl.(Array.length cl - 1) | _ -> anomaly "last_arg" @@ -355,9 +339,7 @@ let general_elim_then_using let elimination_then_using tac predicate (indbindings,elimbindings) c gl = let (ity,path_name,_,t) = reduce_to_ind_goal gl (pf_type_of gl c) in - let elim = - lookup_eliminator (pf_hyps gl) path_name (suff gl (pf_concl gl)) - in + let elim = lookup_eliminator (pf_env gl) path_name (sort_of_goal gl) in general_elim_then_using elim elim_sign tac predicate (indbindings,elimbindings) c gl @@ -409,7 +391,7 @@ let make_elim_branch_assumptions ba gl = | (_, _) -> error "make_elim_branch_assumptions" in makerec ([],[],[],[],[]) ba.branchsign - (try list_firstn ba.nassums (ids_of_sign (pf_untyped_hyps gl)) + (try list_firstn ba.nassums (pf_ids_of_hyps gl) with Failure _ -> anomaly "make_elim_branch_assumptions") let elim_on_ba tac ba gl = tac (make_elim_branch_assumptions ba gl) gl @@ -437,7 +419,7 @@ let make_case_branch_assumptions ba gl = | (_, _) -> error "make_case_branch_assumptions" in makerec ([],[],[],[]) ba.branchsign - (try list_firstn ba.nassums (ids_of_sign (pf_untyped_hyps gl)) + (try list_firstn ba.nassums (pf_ids_of_hyps gl) with Failure _ -> anomaly "make_case_branch_assumptions") let case_on_ba tac ba gl = tac (make_case_branch_assumptions ba gl) gl diff --git a/tactics/tacticals.mli b/tactics/tacticals.mli index 056ac353f..0cdedd8c6 100644 --- a/tactics/tacticals.mli +++ b/tactics/tacticals.mli @@ -34,7 +34,7 @@ val tclWEAK_PROGRESS : tactic -> tactic val tclNTH_HYP : int -> (constr -> tactic) -> tactic val tclMAP : ('a -> tactic) -> 'a list -> tactic val tclLAST_HYP : (constr -> tactic) -> tactic -val tclTRY_sign : (constr -> tactic) -> constr signature -> tactic +val tclTRY_sign : (constr -> tactic) -> var_context -> tactic val tclTRY_HYPS : (constr -> tactic) -> tactic (*val dyn_tclIDTAC : tactic_arg list -> tactic @@ -103,9 +103,7 @@ type branch_assumptions = { indargs : identifier list} (* the inductive arguments *) val sort_of_goal : goal sigma -> sorts -val suff : goal sigma -> constr -> string -val lookup_eliminator : - typed_type signature -> section_path -> string -> constr +(*val suff : goal sigma -> constr -> string*) val general_elim_then_using : constr -> (inductive -> int -> bool list) -> diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 9a03eca97..0beab3f7a 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -11,6 +11,7 @@ open Term open Inductive open Reduction open Environ +open Declare open Evd open Pfedit open Tacred @@ -153,7 +154,7 @@ let reduct_in_concl redfun gl = convert_concl (pf_reduce redfun gl (pf_concl gl)) gl let reduct_in_hyp redfun id gl = - let ty = pf_get_hyp gl id in + let ty = pf_get_hyp_typ gl id in let redfun' = under_casts (fun _ _ -> pf_reduce redfun gl) in convert_hyp id (redfun' (pf_env gl) (project gl) ty) gl @@ -261,7 +262,7 @@ let next_global_ident_away id avoid = next_global_ident_from (lift_ident id) avoid let fresh_id avoid id gl = - next_global_ident_away id (avoid@(ids_of_sign (pf_untyped_hyps gl))) + next_global_ident_away id (avoid@(pf_ids_of_hyps gl)) let id_of_name_with_default s = function | Anonymous -> id_of_string s @@ -407,23 +408,33 @@ let rec intros_move = function tclTHEN (central_intro (IMustBe hyp) destopt false) (intros_move rest) +let occur_var_in_decl id (c,t) = + match c with + | None -> occur_var id (body_of_type t) + | Some body -> occur_var id body || occur_var id (body_of_type t) + +let dependent_in_decl a (c,t) = + match c with + | None -> dependent a (body_of_type t) + | Some body -> dependent a body || dependent a (body_of_type t) + let move_to_rhyp rhyp gl = - let rec get_lhyp lastfixed deptyp = function + let rec get_lhyp lastfixed depdecls = function | [] -> (match rhyp with | None -> lastfixed | Some h -> anomaly ("Hypothesis should occur: "^ (string_of_id h))) - | (hyp,typ) as ht :: rest -> + | (hyp,c,typ) as ht :: rest -> if Some hyp = rhyp then lastfixed - else if List.exists (occur_var hyp) deptyp then - get_lhyp lastfixed (typ::deptyp) rest + else if List.exists (occur_var_in_decl hyp) depdecls then + get_lhyp lastfixed ((c,typ)::depdecls) rest else - get_lhyp (Some hyp) deptyp rest + get_lhyp (Some hyp) depdecls rest in - let sign = pf_untyped_hyps gl in - let (hyp,typ) = hd_sign sign in - match get_lhyp None [typ] (list_of_sign sign) with + let sign = pf_hyps gl in + let (hyp,c,typ) = List.hd sign in + match get_lhyp None [c,typ] sign with | None -> tclIDTAC gl | Some hypto -> move_hyp true hyp hypto gl @@ -585,25 +596,30 @@ let generalize_goal gl c cl = prod_name (Global.env()) (Anonymous, t, cl') let generalize_dep c gl = - let sign = pf_untyped_hyps gl in - let init_ids = ids_of_sign (Global.var_context()) in - let rec seek ((hl,tl) as toquant) h t = - if List.exists (fun id -> occur_var id t) hl or dependent c t then - (h::hl,t::tl) + let sign = pf_hyps gl in + let init_ids = ids_of_var_context (Global.var_context()) in + let rec seek toquant (h,body,t as d) = + if List.exists (fun (id,_,_) -> occur_var_in_decl id (body,t)) toquant + or dependent_in_decl c (body,t) then + d::toquant else toquant in - let (hl,tl) = it_sign seek ([],[]) sign in - let tothin = List.filter (fun id -> not (List.mem id init_ids)) hl in + let to_quantify = List.fold_left seek [] sign in + let qhyps = List.map (fun (id,_,_) -> id) to_quantify in + let tothin = List.filter (fun id -> not (List.mem id init_ids)) qhyps in let tothin' = match c with - | VAR id when mem_sign sign id & not (List.mem id init_ids) -> id::tothin + | VAR id when mem_var_context id sign & not (List.mem id init_ids) + -> id::tothin | _ -> tothin in - let cl' = List.fold_right2 mkNamedProd hl tl (pf_concl gl) in + let cl' = + List.fold_left + (fun c d -> mkNamedProd_or_LetIn d c) (pf_concl gl) to_quantify in let cl'' = generalize_goal gl c cl' in tclTHEN - (apply_type cl'' (c::(List.map (fun id -> VAR id) hl))) + (apply_type cl'' (c::(List.map (fun id -> VAR id) qhyps))) (thin (List.rev tothin')) gl @@ -630,30 +646,42 @@ let quantify lconstr = tclIDTAC *) -(* A dependent cut rule à la sequent calculus *) +(* A dependent cut rule à la sequent calculus + ------------------------------------------ + Sera simplifiable le jour où il y aura un let in primitif dans constr -(* Sera simplifiable le jour où il y aura un let in primitif dans constr *) + [letin_tac b na c occ_ccl occ_hyps gl] transforms + [...x1:T1(c),...,x2:T2(c),... |- G(c)] into + [...x:T;x1:T1(x),...,x2:T2(x),... |- G(x)] if [b] is false or + [...x:T;H:x=c;x1:T1(x),...,x2:T2(x),... |- G(x)] if [b] is true -(* if [occhypl] is empty, [c] is substituted in every hyp where it occurs *) -(* if name = Anonymous, the name is build from the first letter of the type *) + [occ_hyps] and [occ_ccl] tells which occurrences of [c] have to be substd; + if [occ_hyps] is empty, [c] is substituted in every hyp where it occurs; + if name = Anonymous, the name is build from the first letter of the type; -let letin_abstract id c occ_ccl occhypl gl = - let allhyp = occhypl=[] in - let hyps = pf_untyped_hyps gl in - let abstract ((dephyps,deptyps,marks,occl as accu),lhyp) hyp typ = + The tactic first quantify the goal over x1, x2,... then substitute then + re-intro x1, x2,... at their initial place ([marks] is internally + used to remember the place of x1, x2, ...: it is the list of hypotheses on + the left of each x1, ...). +*) +(* +let letin_abstract id c occ_ccl occ_hyps gl = + let allhyp = occ_hyps=[] in + let hyps = pf_hyps gl in + let abstract ((depdecls,marks,occl as accu),lhyp) (hyp,bodyopt,typ as d) = try let occ = if allhyp then [] else List.assoc hyp occl in - let newtyp = subst1 (VAR id) (subst_term_occ occ c typ) in + let newdecl = subst1 (VAR id) (subst_term_occ_decl occ c d) in let newoccl = list_except_assoc hyp occl in - if typ=newtyp then + if d=newdecl then (accu,Some hyp) else - ((hyp::dephyps,newtyp::deptyps,(hyp,lhyp)::marks,newoccl),lhyp) + ((newdecl::depdecls,(hyp,lhyp)::marks,newoccl),lhyp) with Not_found -> (accu,Some hyp) in - let (dephyps,deptyps,marks,rest),_ = - it_sign abstract (([],[],[],occhypl),None) hyps in + let (depdecls,marks,rest),_ = + it_sign abstract (([],[],[],occ_hyps),None) hyps in if rest <> [] then begin let id = fst (List.hd rest) in if mem_sign hyps id @@ -664,14 +692,14 @@ let letin_abstract id c occ_ccl occhypl gl = | None -> (pf_concl gl) | Some occ -> subst1 (VAR id) (subst_term_occ occ c (pf_concl gl)) in - (dephyps,deptyps,marks,ccl) + (depdecls,marks,ccl) -let letin_tac with_eq name c occ_ccl occhypl gl = +let letin_tac with_eq name c occ_ccl occ_hyps gl = let x = id_of_name_using_hdchar (Global.env()) (pf_type_of gl c) name in let hyps = pf_untyped_hyps gl in let id = next_global_ident_away x (ids_of_sign hyps) in if mem_sign hyps id then error "New variable is already declared"; - let (dephyps,deptyps,marks,ccl)= letin_abstract id c occ_ccl occhypl gl in + let (depdecls,marks,ccl)= letin_abstract id c occ_ccl occ_hyps gl in let t = pf_type_of gl c in let (eqc,reflc) = let sort = pf_type_of gl t in @@ -682,8 +710,8 @@ let letin_tac with_eq name c occ_ccl occhypl gl = else error "Not possible with proofs yet" in let heq = next_global_ident_away (id_of_string "Heq") (ids_of_sign hyps) in - let tmpcl = List.fold_right2 mkNamedProd dephyps deptyps ccl in - let tmpargs = List.map (fun id -> VAR id) dephyps in + let tmpcl = List.fold_right mkNamedProd_or_LetIn depdecls ccl in + let tmpargs = List.map (fun (id,_,_) -> VAR id) depdecls in let newcl,args = if with_eq then let eq = applist (eqc,[t;VAR id;c]) in @@ -706,7 +734,8 @@ let dyn_lettac args gl = match args with | [Identifier id; Constr c; Letpatterns (o,l)] -> letin_tac true (Name id) c o l gl | l -> bad_tactic_args "letin" l - +*) +let dyn_lettac a = failwith "TO REDO" (********************************************************************) (* Exact tactics *) @@ -738,13 +767,13 @@ let dyn_exact cc gl = match cc with let (assumption : tactic) = fun gl -> let concl = pf_concl gl in - let rec arec sign = - if isnull_sign sign then error "No such assumption"; - let (s,a) = hd_sign sign in - if pf_conv_x_leq gl a concl then refine (VAR(s)) gl - else arec (tl_sign sign) + let rec arec = function + | [] -> error "No such assumption" + | (id,c,t)::rest -> + if pf_conv_x_leq gl (body_of_type t) concl then refine (VAR id) gl + else arec rest in - arec (pf_untyped_hyps gl) + arec (pf_hyps gl) let dyn_assumption = function | [] -> assumption @@ -966,9 +995,7 @@ let general_elim (c,lbindc) (elimc,lbindelimc) gl = let default_elim (c,lbindc) gl = let (path_name,_,t) = reduce_to_ind (pf_env gl) (project gl) (pf_type_of gl c) in - let elimc = - lookup_eliminator (pf_hyps gl) path_name (suff gl (pf_concl gl)) - in + let elimc = lookup_eliminator (pf_env gl) path_name (sort_of_goal gl) in general_elim (c,lbindc) (elimc,[]) gl @@ -1019,8 +1046,8 @@ let rec is_rec_arg indpath t = let induct_discharge indpath statuslists cname destopt avoid (_,t) = let (lstatus,rstatus) = statuslists in let tophyp = ref None in - let (l,_) = decompose_prod t in - let n = List.length (List.filter (fun (_,t') -> is_rec_arg indpath t') l) in + let (l,_) = decompose_prod_assum t in + let n = List.length (List.filter (fun (_,_,t') -> is_rec_arg indpath (body_of_type t')) l) in let recvarname = if n=1 then cname @@ -1057,14 +1084,9 @@ let induct_discharge indpath statuslists cname destopt avoid (_,t) = substitutions aussi sur l'argument voisin *) (* Marche pas... faut prendre en compte l'occurence précise... *) - +(* let atomize_param_of_ind hyp0 gl = - let tmptyp0 = - try - (snd(lookup_sign hyp0 (pf_untyped_hyps gl))) - with Not_found -> - error ("No such hypothesis : " ^ (string_of_id hyp0)) - in + let tmptyp0 = pf_get_hyp_typ gl hyp0 in let (mind,indtyp,typ0) = pf_reduce_to_mind gl tmptyp0 in let mis = Global.lookup_mind_specif mind in let nparams = mis_nparams mis in @@ -1073,7 +1095,7 @@ let atomize_param_of_ind hyp0 gl = (* le gl est important pour ne pas préévaluer *) let rec atomize_one i avoid gl = if i<>nparams then - let tmptyp0 = pf_get_hyp gl hyp0 in + let tmptyp0 = pf_get_hyp_typ gl hyp0 in let (_,indtyp,_) = pf_reduce_to_mind gl tmptyp0 in match (destAppL (whd_castapp indtyp)).(i) with | VAR id when not (List.exists (occur_var id) avoid) -> @@ -1309,7 +1331,7 @@ let new_induct c gl = tclTHEN (letin_tac true (Name id) c (Some []) []) (induction_with_atomization_of_ind_arg id) gl - +*) (***) (* The registered tactic, which calls the default elimination @@ -1330,10 +1352,14 @@ let dyn_elim = function | l -> bad_tactic_args "elim" l (* Induction tactics *) - +(* let induct s = tclORELSE (tclTHEN (intros_until s) (tclLAST_HYP simplest_elim)) (induction_from_context s) +*) +let induct s = + tclTHEN (intros_until s) (tclLAST_HYP simplest_elim) + let induct_nodep n = tclTHEN (intros_do n) (tclLAST_HYP simplest_elim) @@ -1398,9 +1424,7 @@ let elim_scheme_type elim t gl = let elim_type t gl = let (path_name,tind,t) = reduce_to_ind (pf_env gl) (project gl) t in - let elimc = - lookup_eliminator (pf_hyps gl) path_name (suff gl (pf_concl gl)) - in + let elimc = lookup_eliminator (pf_env gl) path_name (sort_of_goal gl) in match t with | DOP2(Prod,_,_) -> error "Not an inductive definition" | _ -> elim_scheme_type elimc tind gl @@ -1442,7 +1466,7 @@ let dyn_case_type = function *) let andE id gl = - let t = pf_get_hyp gl id in + let t = pf_get_hyp_typ gl id in if is_conjunction (pf_hnf_constr gl t) then (tclTHEN (simplest_elim (VAR id)) (tclDO 2 intro)) gl else @@ -1455,7 +1479,7 @@ let dAnd cls gl = | Some id -> andE id gl let orE id gl = - let t = pf_get_hyp gl id in + let t = pf_get_hyp_typ gl id in if is_disjunction (pf_hnf_constr gl t) then (tclTHEN (simplest_elim (VAR id)) intro) gl else @@ -1468,7 +1492,7 @@ let dorE b cls gl = | None -> (if b then right else left) [] gl let impE id gl = - let t = pf_get_hyp gl id in + let t = pf_get_hyp_typ gl id in if is_imp_term (pf_hnf_constr gl t) then let (dom, _, rng) = destProd (pf_hnf_constr gl t) in (tclTHENS (cut_intro rng) @@ -1490,7 +1514,7 @@ let dImp cls gl = (* Contradiction *) let contradiction_on_hyp id gl = - let hyp = pf_get_hyp gl id in + let hyp = pf_get_hyp_typ gl id in if is_empty_type hyp then simplest_elim (VAR id) gl else @@ -1503,10 +1527,10 @@ let absurd c gls = (tclTHEN (elim_type falseterm) (cut c)) ([(tclTHENS (cut (applist(pf_global gls (id_of_string "not"),[c]))) - ([(tclTHEN (intros) + ([(tclTHEN intros ((fun gl -> - let (ida,_) = pf_nth_hyp gl 1 - and (idna,_) = pf_nth_hyp gl 2 in + let ida = pf_nth_hyp_id gl 1 + and idna = pf_nth_hyp_id gl 2 in exact (applist(VAR idna,[VAR ida])) gl))); tclIDTAC])); tclIDTAC])) gls @@ -1626,32 +1650,33 @@ let dyn_transitivity = function let abstract_subproof name tac gls = let env = Global.env() in let current_sign = Global.var_context() - and global_sign = pf_untyped_hyps gls in - let sign = Sign.sign_it - (fun id typ s -> - if mem_sign current_sign id then s else add_sign (id,typ) s) - global_sign nil_sign + and global_sign = pf_hyps gls in + let sign = List.fold_right + (fun (id,_,_ as d) s -> + if mem_var_context id current_sign then s else add_var d s) + global_sign empty_var_context in let na = next_global_ident_away name - (ids_of_sign global_sign) in - let concl = Sign.it_sign (fun t id typ -> mkNamedProd id typ t) + (ids_of_var_context global_sign) in + let concl = List.fold_left (fun t d -> mkNamedProd_or_LetIn d t) (pf_concl gls) sign in let env' = change_hyps (fun _ -> current_sign) env in let lemme = start_proof na Declare.NeverDischarge env' concl; let _,(const,strength) = try - by (tclCOMPLETE (tclTHEN (tclDO (sign_length sign) intro) tac)); - release_proof () + by (tclCOMPLETE (tclTHEN (tclDO (List.length sign) intro) tac)); + let r = cook_proof () in + delete_current_proof (); r with e when catchable_exception e -> - (abort_current_proof(); raise e) + (delete_current_proof(); raise e) in (* Faudrait un peu fonctionnaliser cela *) Declare.declare_constant na (const,strength); let newenv = Global.env() in Declare.construct_reference newenv CCI na in exact (applist (lemme, - List.map (fun id -> VAR id) (List.rev (ids_of_sign sign)))) + List.map (fun id -> VAR id) (List.rev (ids_of_var_context sign)))) gls let tclABSTRACT name_op tac gls = diff --git a/tactics/tauto.ml b/tactics/tauto.ml index 0130566b4..47f7a40cf 100644 --- a/tactics/tauto.ml +++ b/tactics/tauto.ml @@ -23,17 +23,19 @@ open Proof_trees open Clenv open Pattern +(* Faut-il comparer le corps des définitions de l'environnement ? *) let hlset_subset hls1 hls2 = - List.for_all (fun e -> List.exists (fun e' -> eq_constr e e') hls2) hls1 + List.for_all + (fun (_,_,e) -> List.exists + (fun (_,_,e') -> eq_constr (body_of_type e) (body_of_type e')) + hls2) + hls1 let hlset_eq hls1 hls2 = hlset_subset hls1 hls2 & hlset_subset hls2 hls1 let eq_gls g1 g2 = - eq_constr (pf_concl g1) (pf_concl g2) - & (let hl1 = vals_of_sign (pf_untyped_hyps g1) - and hl2 = vals_of_sign (pf_untyped_hyps g2) in - hlset_eq hl1 hl2) + eq_constr (pf_concl g1) (pf_concl g2) & hlset_eq (pf_hyps g1) (pf_hyps g2) let gls_memb bTS g = List.exists (eq_gls g) bTS @@ -1715,14 +1717,13 @@ let rec lisvarprop = function (*-- Dado el ambiente COQ construye el lado izquierdo de un secuente (hipotesis) --*) let rec constr_lseq gls = function - | ([],[]) -> [] - | (idx::idy,hx::hy) -> - (match (pf_type_of gls hx) with - | DOP0 (Sort (Prop Null)) -> - (TVar(string_of_id idx),tauto_of_cci_fmla gls hx) - :: constr_lseq gls (idy,hy) - |_-> constr_lseq gls (idy,hy)) - | _ -> [] + | [] -> [] + | (idx,c,hx)::rest -> + match Retyping.get_sort_of (pf_env gls) (project gls) (incast_type hx) with + | Prop Null -> + (TVar(string_of_id idx),tauto_of_cci_fmla gls (body_of_type hx)) + :: constr_lseq gls rest + |_-> constr_lseq gls rest (*-- Dado un estado COQ construye el lado derecho de un secuente (conclusion) --*) @@ -1735,7 +1736,7 @@ let rec pos_lis x = function | a::r -> if (x=a) then 1 else 1 + (pos_lis x r) (*-- Construye un termino constr dado una formula tauto --*) -let rec cci_of_tauto_fml env = +let rec cci_of_tauto_fml () = let cAnd = global_reference CCI (id_of_string "and") and cOr = global_reference CCI (id_of_string "or") and cFalse = global_reference CCI (id_of_string "False") @@ -1743,13 +1744,13 @@ let rec cci_of_tauto_fml env = and cEq = global_reference CCI (id_of_string "eq") in function | FAnd(a,b) -> applistc cAnd - [cci_of_tauto_fml env a;cci_of_tauto_fml env b] + [cci_of_tauto_fml () a;cci_of_tauto_fml () b] | FOr(a,b) -> applistc cOr - [cci_of_tauto_fml env a; cci_of_tauto_fml env b] + [cci_of_tauto_fml () a; cci_of_tauto_fml () b] | FEq(a,b,c)-> applistc cEq - [cci_of_tauto_fml env a; cci_of_tauto_fml env b; - cci_of_tauto_fml env c] - | FImp(a,b) -> mkArrow (cci_of_tauto_fml env a) (cci_of_tauto_fml env b) + [cci_of_tauto_fml () a; cci_of_tauto_fml () b; + cci_of_tauto_fml () c] + | FImp(a,b) -> mkArrow (cci_of_tauto_fml () a) (cci_of_tauto_fml () b) | FPred a -> a | FFalse -> cFalse | FTrue -> cTrue @@ -1761,10 +1762,11 @@ let rec cci_of_tauto_fml env = let search env id = try - (match lookup_id id (Environ.context env) with - | RELNAME (n,_) -> Rel n - | GLOBNAME _ -> VAR id) + Rel (fst (lookup_rel_id id (Environ.rel_context env))) with Not_found -> + if mem_var_context id (Environ.var_context env) then + VAR id + else global_reference CCI id (*-- Construye un termino constr de un termino tauto --*) @@ -1786,37 +1788,37 @@ let cci_of_tauto_term env t = search env (id_of_string x)) with _ -> raise TacticFailure) | TZ(f,x) -> applistc cFalse_ind - [cci_of_tauto_fml env f;ter_constr l x] + [cci_of_tauto_fml () f;ter_constr l x] | TSum(t1,t2) -> ter_constr l t1 | TExi (x) -> (try search env (id_of_string x) with _ -> raise TacticFailure) | TApl(_,_,t1,t2) -> applistc (ter_constr l t1) [ter_constr l t2] | TPar(f1,f2,t1,t2) -> applistc cconj - [cci_of_tauto_fml env f1;cci_of_tauto_fml env f2; + [cci_of_tauto_fml () f1;cci_of_tauto_fml () f2; ter_constr l t1;ter_constr l t2] | TInl(f1,f2,t) -> applistc cor_introl - [cci_of_tauto_fml env f1;cci_of_tauto_fml env f2; + [cci_of_tauto_fml () f1;cci_of_tauto_fml () f2; ter_constr l t] | TInr(f1,f2,t) -> applistc cor_intror - [cci_of_tauto_fml env f1;cci_of_tauto_fml env f2; + [cci_of_tauto_fml () f1;cci_of_tauto_fml () f2; ter_constr l t] | TFst(f1,f2,t) -> applistc cproj1 - [cci_of_tauto_fml env f1;cci_of_tauto_fml env f2; + [cci_of_tauto_fml () f1;cci_of_tauto_fml () f2; ter_constr l t] | TSnd(f1,f2,t) -> applistc cproj2 - [cci_of_tauto_fml env f1;cci_of_tauto_fml env f2; + [cci_of_tauto_fml () f1;cci_of_tauto_fml () f2; ter_constr l t] | TRefl(f1,f2) -> applistc crefl - [cci_of_tauto_fml env f1;cci_of_tauto_fml env f2] + [cci_of_tauto_fml () f1;cci_of_tauto_fml () f2] | TSim(f1,f2,f3,t) -> applistc csim - [cci_of_tauto_fml env f1;cci_of_tauto_fml env f2; - cci_of_tauto_fml env f3;ter_constr l t] + [cci_of_tauto_fml () f1;cci_of_tauto_fml () f2; + cci_of_tauto_fml () f3;ter_constr l t] | TCase(lf,lt) -> applistc cor_ind - ((List.map (cci_of_tauto_fml env) lf)@ + ((List.map (cci_of_tauto_fml ()) lf)@ (List.map (ter_constr l) lt)) | TFun(n,f,t) -> Environ.lambda_name env - (Name(id_of_string n ), cci_of_tauto_fml env f,ter_constr (n::l) t) + (Name(id_of_string n ), cci_of_tauto_fml () f,ter_constr (n::l) t) | TTrue -> ci | TLis _ -> raise TacticFailure | TLister _ -> raise TacticFailure @@ -1843,17 +1845,15 @@ let exacto tt gls = let subbuts l hyp = cut_in_parallelUsing (lisvar l) - (List.map (cci_of_tauto_fml (gLOB hyp)) (lisfor l)) + (List.map (cci_of_tauto_fml ()) (lisfor l)) let t_exacto = ref (TVar "#") let tautoOR ti gls = - let hyp = pf_untyped_hyps gls in let thyp = pf_hyps gls in - hipvar := ids_of_sign hyp; + hipvar := ids_of_var_context thyp; let but = pf_concl gls in - let seq = (constr_lseq gls (ids_of_sign hyp,vals_of_sign hyp), - constr_rseq gls but) in + let seq = (constr_lseq gls thyp, constr_rseq gls but) in let vp = lisvarprop (ante seq) in match naive ti vp seq with | {arb=Nil} -> diff --git a/tactics/wcclausenv.ml b/tactics/wcclausenv.ml index 48ae03cc9..005be137e 100644 --- a/tactics/wcclausenv.ml +++ b/tactics/wcclausenv.ml @@ -22,10 +22,10 @@ open Clenv type wc = walking_constraints let pf_get_new_id id gls = - next_ident_away id (ids_of_sign (pf_untyped_hyps gls)) + next_ident_away id (pf_ids_of_hyps gls) let pf_get_new_ids ids gls = - let avoid = ids_of_sign (pf_untyped_hyps gls) in + let avoid = pf_ids_of_hyps gls in List.fold_right (fun id acc -> (next_ident_away id (acc@avoid))::acc) ids [] @@ -93,7 +93,7 @@ let clenv_constrain_with_bindings bl clause = let add_prod_rel sigma (t,env) = match t with | DOP2(Prod,c1,(DLAM(na,b))) -> - (b,push_rel (na,Typing.execute_type env sigma c1) env) + (b,push_rel_decl (na,Retyping.get_assumption_of env sigma c1) env) | _ -> failwith "add_prod_rel" let rec add_prods_rel sigma (t,env) = diff --git a/toplevel/command.ml b/toplevel/command.ml index 08c740013..60d9e5f13 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -91,8 +91,9 @@ let hypothesis_def_var is_refining ident n c = parameter_def_var ident c | DischargeAt disch_sp -> if Lib.is_section_p disch_sp then begin + let t = interp_type Evd.empty (Global.env()) c in declare_variable (id_of_string ident) - (interp_type Evd.empty (Global.env()) c,n,false); + (SectionLocalDecl t,n,false); if is_verbose() then message (ident ^ " is assumed"); if is_refining then mSGERRNL [< 'sTR"Warning: Variable "; 'sTR ident; @@ -145,7 +146,7 @@ let build_mutual lparams lnamearconstrs finite = (fun (env,arl) (recname,arityc,_) -> let arity = typed_type_of_com Evd.empty env (mkProdCit lparams arityc) in - let env' = Environ.push_rel (Name recname,arity) env in + let env' = Environ.push_rel_decl (Name recname,arity) env in (env', (arity::arl))) (env0,[]) lnamearconstrs in @@ -218,8 +219,8 @@ let build_recursive lnameargsardef = List.fold_left (fun (env,arl) (recname,lparams,arityc,_) -> let arity = typed_type_of_com Evd.empty env (mkProdCit lparams arityc) in - declare_variable recname (body_of_type arity,NeverDischarge,false); - (Environ.push_var (recname,arity) env, (arity::arl))) + declare_variable recname (SectionLocalDecl (body_of_type arity),NeverDischarge,false); + (Environ.push_var_decl (recname,arity) env, (arity::arl))) (env0,[]) lnameargsardef with e -> States.unfreeze fs; raise e @@ -281,8 +282,8 @@ let build_corecursive lnameardef = List.fold_left (fun (env,arl) (recname,arityc,_) -> let arity = typed_type_of_com Evd.empty env0 arityc in - declare_variable recname (body_of_type arity,NeverDischarge,false); - (Environ.push_var (recname,arity) env, (arity::arl))) + declare_variable recname (SectionLocalDecl (body_of_type arity),NeverDischarge,false); + (Environ.push_var_decl (recname,arity) env, (arity::arl))) (env0,[]) lnameardef with e -> States.unfreeze fs; raise e @@ -376,15 +377,17 @@ let start_proof_com sopt stre com = Pfedit.start_proof id stre env (interp_type Evd.empty env com) let save_named opacity = - let id,(const,strength) = Pfedit.release_proof () in + let id,(const,strength) = Pfedit.cook_proof () in declare_constant id (const,strength); + Pfedit.delete_current_proof (); if Options.is_verbose() then message ((string_of_id id) ^ " is defined") let save_anonymous opacity save_ident strength = - let id,(const,_) = Pfedit.release_proof () in + let id,(const,_) = Pfedit.cook_proof () in if atompart_of_id id <> "Unnamed_thm" then message("Overriding name "^(string_of_id id)^" and using "^save_ident); declare_constant (id_of_string save_ident) (const,strength); + Pfedit.delete_current_proof (); if Options.is_verbose() then message (save_ident ^ " is defined") let save_anonymous_thm opacity id = diff --git a/toplevel/discharge.ml b/toplevel/discharge.ml index ebf70569d..91ce37753 100644 --- a/toplevel/discharge.ml +++ b/toplevel/discharge.ml @@ -26,8 +26,7 @@ let recalc_sp sp = let whd_all c = whd_betadeltaiota (Global.env()) Evd.empty c -let generalize_type id var c = - typed_product_without_universes (Name id) var (typed_app (subst_var id) c) +let generalize_type = generalize_without_universes type modification_action = ABSTRACT | ERASE @@ -100,16 +99,16 @@ let under_dlams f = apprec let abstract_inductive ids_to_abs hyps inds = - let abstract_one_var id ty inds = + let abstract_one_var d inds = let ntyp = List.length inds in let new_refs = list_tabulate (fun k -> applist(Rel (k+2),[Rel 1])) ntyp in let inds' = List.map (function (tname,arity,cnames,lc) -> - let arity' = generalize_type id ty arity in + let arity' = generalize_type d arity in let lc' = List.map - (fun b -> generalize_type id ty (typed_app (substl new_refs) b)) + (fun b -> generalize_type d (typed_app (substl new_refs) b)) lc in (tname,arity',cnames,lc')) @@ -118,11 +117,13 @@ let abstract_inductive ids_to_abs hyps inds = (inds',ABSTRACT) in let abstract_once ((hyps,inds,modl) as sofar) id = - if isnull_sign hyps or id <> fst (hd_sign hyps) then - sofar - else - let (inds',modif) = abstract_one_var id (snd (hd_sign hyps)) inds in - (tl_sign hyps,inds',modif::modl) + match hyps with + | [] -> sofar + | (hyp,c,t as d)::rest -> + if id <> hyp then sofar + else + let (inds',modif) = abstract_one_var d inds in + (rest, inds', modif::modl) in let (_,inds',revmodl) = List.fold_left abstract_once (hyps,inds,[]) ids_to_abs in @@ -134,22 +135,20 @@ let abstract_inductive ids_to_abs hyps inds = let abstract_constant ids_to_abs hyps (body,typ) = let abstract_once ((hyps,body,typ,modl) as sofar) id = - if isnull_sign hyps or id <> fst(hd_sign hyps) then - sofar + match hyps with + | [] -> sofar + | (hyp,c,t as decl)::rest -> + if hyp <> id then sofar else - let name = Name id in - let var = snd (hd_sign hyps) in - let cvar = incast_type var in let body' = match body with | None -> None | Some { contents = Cooked c } -> - Some (ref (Cooked (mkLambda name cvar (subst_var id c)))) + Some (ref (Cooked (mkNamedLambda_or_LetIn decl c))) | Some { contents = Recipe f } -> - Some (ref (Recipe - (fun () -> mkLambda name cvar (subst_var id (f()))))) + Some (ref (Recipe (fun () -> mkNamedLambda_or_LetIn decl (f())))) in - let typ' = generalize_type id var typ in - (tl_sign hyps,body',typ',ABSTRACT::modl) + let typ' = generalize_type decl typ in + (rest, body', typ', ABSTRACT::modl) in let (_,body',typ',revmodl) = List.fold_left abstract_once (hyps,body,typ,[]) ids_to_abs in @@ -244,7 +243,7 @@ let process_constant osecsp nsecsp oldenv (ids_to_discard,modlist) cb = let body = expmod_constant_value cb.const_opaque oldenv modlist cb.const_body in let typ = expmod_type oldenv modlist cb.const_type in - let hyps = map_sign_typ (expmod_type oldenv modlist) cb.const_hyps in + let hyps = map_var_context (expmod_constr oldenv modlist) cb.const_hyps in let (body',typ',modl) = abstract_constant ids_to_discard hyps (body,typ) in let mods = [ (Const osecsp, DO_ABSTRACT(Const nsecsp,modl)) ] in (body', body_of_type typ', mods) @@ -267,7 +266,7 @@ let inductive_message inds = 'sPC; 'sTR "are discharged.">])) type discharge_operation = - | Variable of identifier * constr * strength * bool + | Variable of identifier * section_variable_entry * strength * bool | Parameter of identifier * constr | Constant of identifier * constant_entry * strength | Inductive of mutual_inductive_entry @@ -280,13 +279,20 @@ let process_object oldenv sec_sp (ops,ids_to_discard,work_alist) (sp,lobj) = let tag = object_tag lobj in match tag with | "VARIABLE" -> - let (id,a,stre,sticky) = out_variable sp in + let ((id,c,t),stre,sticky) = out_variable sp in if stre = (DischargeAt sec_sp) or ids_to_discard <> [] then (ops,id::ids_to_discard,work_alist) else - let expmod_a = expmod_constr oldenv work_alist (body_of_type a) in - (Variable (id,expmod_a,stre,sticky) :: ops, - ids_to_discard,work_alist) + let newdecl = + match c with + | None -> + SectionLocalDecl + (expmod_constr oldenv work_alist (body_of_type t)) + | Some body -> + SectionLocalDef + (expmod_constr oldenv work_alist body) + in + (Variable (id,newdecl,stre,sticky) :: ops, ids_to_discard,work_alist) | "CONSTANT" | "PARAMETER" -> let stre = constant_or_parameter_strength sp in diff --git a/toplevel/himsg.ml b/toplevel/himsg.ml index 413aafce6..9608f5af5 100644 --- a/toplevel/himsg.ml +++ b/toplevel/himsg.ml @@ -22,15 +22,16 @@ let guill s = "\""^s^"\"" let explain_unbound_rel k ctx n = let ctx = make_all_name_different ctx in - let pe = pr_ne_env [< 'sTR"in environment" >] k ctx in + let pe = pr_ne_context_of [< 'sTR"in environment" >] k ctx in [< 'sTR"Unbound reference: "; pe; 'fNL; 'sTR"The reference "; 'iNT n; 'sTR" is free" >] let explain_not_type k ctx j = let ctx = make_all_name_different ctx in - let pe = pr_ne_env [< 'sTR"In environment" >] k ctx in + let pe = pr_ne_context_of [< 'sTR"In environment" >] k ctx in let pc,pt = prjudge_env ctx j in - [< pe; 'cUT; 'sTR "the term"; 'bRK(1,1); pc; 'sPC; 'sTR"has type"; 'sPC; + [< pe; 'fNL; 'sTR "the term"; 'bRK(1,1); pc; 'sPC; + 'sTR"has type"; 'sPC; pt; 'sPC; 'sTR"which should be Set, Prop or Type." >];; let explain_bad_assumption k ctx c = @@ -92,9 +93,9 @@ let explain_ill_formed_branch k ctx c i actty expty = let explain_generalization k ctx (name,var) j = let ctx = make_all_name_different ctx in - let pe = pr_ne_env [< 'sTR"in environment" >] k ctx in + let pe = pr_ne_context_of [< 'sTR"in environment" >] k ctx in let pv = prtype_env ctx var in - let (pc,pt) = prjudge_env (add_rel (name,var) ctx) j in + let (pc,pt) = prjudge_env (push_rel_decl (name,var) ctx) j in [< 'sTR"Illegal generalization: "; pe; 'fNL; 'sTR"Cannot generalize"; 'bRK(1,1); pv; 'sPC; 'sTR"over"; 'bRK(1,1); pc; 'sTR","; 'sPC; 'sTR"it has type"; 'sPC; pt; @@ -102,7 +103,7 @@ let explain_generalization k ctx (name,var) j = let explain_actual_type k ctx c ct pt = let ctx = make_all_name_different ctx in - let pe = pr_ne_env [< 'sTR"In environment" >] k ctx in + let pe = pr_ne_context_of [< 'sTR"In environment" >] k ctx in let pc = prterm_env ctx c in let pct = prterm_env ctx ct in let pt = prterm_env ctx pt in @@ -113,10 +114,14 @@ let explain_actual_type k ctx c ct pt = let explain_cant_apply_bad_type k ctx (n,exptyp,actualtyp) rator randl = let ctx = make_all_name_different ctx in -(* let pe = pr_ne_env [< 'sTR"in environment" >] k ctx in*) +(* let pe = pr_ne_context_of [< 'sTR"in environment" >] k ctx in*) let pr,prt = prjudge_env ctx rator in - let term_string = if List.length randl > 1 then "terms" else "term" in - let many = match n mod 10 with 1 -> "st" | 2 -> "nd" | _ -> "th" in + let term_string1,term_string2 = + if List.length randl > 1 then + let many = match n mod 10 with 1 -> "st" | 2 -> "nd" | _ -> "th" in + "terms", "The "^(string_of_int n)^many^" term" + else + "term","This term" in let appl = prlist_with_sep pr_fnl (fun c -> let pc,pct = prjudge_env ctx c in @@ -125,15 +130,14 @@ let explain_cant_apply_bad_type k ctx (n,exptyp,actualtyp) rator randl = [< 'sTR"Illegal application (Type Error): "; (* pe; *) 'fNL; 'sTR"The term"; 'bRK(1,1); pr; 'sPC; 'sTR"of type"; 'bRK(1,1); prt; 'sPC ; - 'sTR("cannot be applied to the "^term_string); 'fNL; + 'sTR("cannot be applied to the "^term_string1); 'fNL; 'sTR" "; v 0 appl; 'fNL; - 'sTR"The ";'iNT n; 'sTR (many^" term of type"); 'bRK(1,1); - prterm_env ctx actualtyp; 'sPC; - 'sTR"should be coercible to"; 'bRK(1,1); prterm_env ctx exptyp >] + 'sTR (term_string2^" has type"); 'bRK(1,1); prterm_env ctx exptyp; 'sPC; + 'sTR"which should be coercible to"; 'bRK(1,1); prterm_env ctx actualtyp >] let explain_cant_apply_not_functional k ctx rator randl = let ctx = make_all_name_different ctx in -(* let pe = pr_ne_env [< 'sTR"in environment" >] k ctx in*) +(* let pe = pr_ne_context_of [< 'sTR"in environment" >] k ctx in*) let pr = prterm_env ctx rator.uj_val in let prt = prterm_env ctx (body_of_type rator.uj_type) in let term_string = if List.length randl > 1 then "terms" else "term" in @@ -356,22 +360,19 @@ let explain_refiner_error = function | NotWellTyped c -> explain_refiner_not_well_typed c | BadTacticArgs (s,l) -> explain_refiner_bad_tactic_args s l -let error_non_strictly_positive k lna c v = - let env = assumptions_for_print lna in +let error_non_strictly_positive k env c v = let pc = prterm_env env c in let pv = prterm_env env v in [< 'sTR "Non strictly positive occurrence of "; pv; 'sTR " in"; 'bRK(1,1); pc >] -let error_ill_formed_inductive k lna c v = - let env = assumptions_for_print lna in +let error_ill_formed_inductive k env c v = let pc = prterm_env env c in let pv = prterm_env env v in [< 'sTR "Not enough arguments applied to the "; pv; 'sTR " in"; 'bRK(1,1); pc >] -let error_ill_formed_constructor k lna c v = - let env = assumptions_for_print lna in +let error_ill_formed_constructor k env c v = let pc = prterm_env env c in let pv = prterm_env env v in [< 'sTR "The conclusion of"; 'bRK(1,1); pc; 'bRK(1,1); @@ -385,8 +386,7 @@ let str_of_nth n = | 3 -> "rd" | _ -> "th") -let error_bad_ind_parameters k lna c n v1 v2 = - let env = assumptions_for_print lna in +let error_bad_ind_parameters k env c n v1 v2 = let pc = prterm_env_at_top env c in let pv1 = prterm_env env v1 in let pv2 = prterm_env env v2 in @@ -412,7 +412,7 @@ let error_not_allowed_case_analysis dep kind i = [< 'sTR (if dep then "Dependent" else "Non Dependent"); 'sTR " case analysis on sort: "; print_sort kind; 'fNL; 'sTR "is not allowed for inductive definition: "; - pr_inductive (Global.context()) i >] + pr_inductive (Global.env()) i >] let error_bad_induction dep indid kind = [<'sTR (if dep then "Dependent" else "Non dependend"); @@ -425,10 +425,10 @@ let error_not_mutual_in_scheme () = let explain_inductive_error = function (* These are errors related to inductive constructions *) - | NonPos (lna,c,v) -> error_non_strictly_positive CCI lna c v - | NotEnoughArgs (lna,c,v) -> error_ill_formed_inductive CCI lna c v - | NotConstructor (lna,c,v) -> error_ill_formed_constructor CCI lna c v - | NonPar (lna,c,n,v1,v2) -> error_bad_ind_parameters CCI lna c n v1 v2 + | NonPos (env,c,v) -> error_non_strictly_positive CCI env c v + | NotEnoughArgs (env,c,v) -> error_ill_formed_inductive CCI env c v + | NotConstructor (env,c,v) -> error_ill_formed_constructor CCI env c v + | NonPar (env,c,n,v1,v2) -> error_bad_ind_parameters CCI env c n v1 v2 | SameNamesTypes id -> error_same_names_types id | SameNamesConstructors (id,cid) -> error_same_names_constructors id cid | NotAnArity id -> error_not_an_arity id diff --git a/toplevel/himsg.mli b/toplevel/himsg.mli index edebaaf52..fabe32634 100644 --- a/toplevel/himsg.mli +++ b/toplevel/himsg.mli @@ -5,14 +5,14 @@ open Pp open Names open Indtypes -open Sign +open Environ open Type_errors open Logic (*i*) (* This module provides functions to explain the type errors. *) -val explain_type_error : path_kind -> context -> type_error -> std_ppcmds +val explain_type_error : path_kind -> env -> type_error -> std_ppcmds val explain_inductive_error : inductive_error -> std_ppcmds diff --git a/toplevel/record.ml b/toplevel/record.ml index 268ebd437..0fc6f85a8 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -72,14 +72,14 @@ let typecheck_params_and_field ps fs = List.fold_left (fun (env,newps) (id,t) -> let tj = typed_type_of_com Evd.empty env t in - (Environ.push_var (id,tj) env,(id,body_of_type tj)::newps)) + (Environ.push_var_decl (id,tj) env,(id,body_of_type tj)::newps)) (env0,[]) ps in let env2,newfs = List.fold_left (fun (env,newfs) (id,t) -> let tj = typed_type_of_com Evd.empty env t in - (Environ.push_var (id,tj) env,(id,body_of_type tj)::newfs)) (env1,[]) fs + (Environ.push_var_decl (id,tj) env,(id,body_of_type tj)::newfs)) (env1,[]) fs in List.rev(newps),List.rev(newfs) diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index d483b1e6a..32a30d144 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -48,13 +48,13 @@ let show_script () = let pts = get_pftreestate () in let pf = proof_of_pftreestate pts and evc = evc_of_pftreestate pts in - mSGNL(Refiner.print_script true (ts_it evc) (Global.var_context()) pf) + mSGNL(Refiner.print_script true evc (Global.var_context()) pf) let show_prooftree () = let pts = get_pftreestate () in let pf = proof_of_pftreestate pts and evc = evc_of_pftreestate pts in - mSG(Refiner.print_proof (ts_it evc) (Global.var_context()) pf) + mSG(Refiner.print_proof evc (Global.var_context()) pf) let show_open_subgoals () = let pfts = get_pftreestate () in @@ -83,7 +83,6 @@ let show_open_subgoals_focused () = let show_node () = let pts = get_pftreestate () in let pf = proof_of_pftreestate pts - and evc = evc_of_pftreestate pts and cursor = cursor_of_pftreestate pts in mSG [< prlist_with_sep pr_spc pr_int (List.rev cursor); 'fNL ; prgl pf.goal ; 'fNL ; @@ -214,7 +213,7 @@ let _ = (* Managing states *) let abort_refine f x = - if Pfedit.refining() then abort_all_proofs (); + if Pfedit.refining() then delete_all_proofs (); f x (* used to be: error "Must save or abort current goal first" *) @@ -352,10 +351,10 @@ let _ = (function | [VARG_IDENTIFIER id] -> (fun () -> - abort_proof id; + delete_proof id; message ("Goal "^(string_of_id id)^" aborted")) | [] -> (fun () -> - abort_current_proof (); + delete_current_proof (); message "Current goal aborted") | _ -> bad_vernac_args "ABORT") @@ -364,7 +363,7 @@ let _ = (function | [] -> (fun () -> if refining() then begin - abort_all_proofs (); + delete_all_proofs (); message "Current goals aborted" end else error "No proof-editing in progress") @@ -529,15 +528,14 @@ let _ = let pts = get_pftreestate () in let pf = proof_of_pftreestate pts in let cursor = cursor_of_pftreestate pts in - let evc = ts_it (evc_of_pftreestate pts) in - let (pfterm,meta_types) = - Refiner.extract_open_proof (var_context pf.goal.evar_env) pf in + let evc = evc_of_pftreestate pts in + let (pfterm,meta_types) = extract_open_pftreestate pts in mSGNL [< 'sTR"LOC: " ; prlist_with_sep pr_spc pr_int (List.rev cursor); 'fNL ; 'sTR"Subgoals" ; 'fNL ; prlist (fun (mv,ty) -> - [< 'iNT mv ; 'sTR" -> " ; prterm ty ; 'fNL >]) + [< 'iNT mv ; 'sTR" -> " ; prtype ty ; 'fNL >]) meta_types; 'sTR"Proof: " ; prterm (nf_ise1 evc pfterm) >]) | _ -> bad_vernac_args "ShowProof") @@ -550,8 +548,7 @@ let _ = let pts = get_pftreestate () in let pf = proof_of_pftreestate pts and cursor = cursor_of_pftreestate pts in - let (pfterm,meta_types) = - Refiner.extract_open_proof (var_context pf.goal.evar_env) pf in + let (pfterm,_) = extract_open_pftreestate pts in let message = try Typeops.control_only_guard pf.goal.evar_env @@ -590,7 +587,7 @@ let _ = | (n::l) -> aux (Tacmach.traverse n pts) l in let pts = aux pts (l@[-1]) in let pf = proof_of_pftreestate pts in - mSG (Refiner.print_script true (ts_it evc) (Global.var_context()) pf)) + mSG (Refiner.print_script true evc (Global.var_context()) pf)) let _ = add "ExplainProofTree" @@ -607,7 +604,7 @@ let _ = | (n::l) -> aux (Tacmach.traverse n pts) l in let pts = aux pts (l@[-1]) in let pf = proof_of_pftreestate pts in - mSG (Refiner.print_proof (ts_it evc) (Global.var_context()) pf)) + mSG (Refiner.print_proof evc (Global.var_context()) pf)) let _ = add "ShowProofs" @@ -719,7 +716,7 @@ let _ = mSGNL [< 'sTR"Error: checking of theorem " ; print_id s ; 'sPC ; 'sTR"failed" ; 'sTR"... converting to Axiom" >]; - abort_proof s; + delete_proof s; parameter_def_var (string_of_id s) com end else errorlabstrm "vernacentries__TheoremProof" |