diff options
author | barras <barras@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2002-02-15 18:02:05 +0000 |
---|---|---|
committer | barras <barras@85f007b7-540e-0410-9357-904b9bb8a0f7> | 2002-02-15 18:02:05 +0000 |
commit | 1e6c3e993fd33d01713aae34a8cefbc210b3898a (patch) | |
tree | 2f8e2aba2c50587146ac4100bb8bf3c426fca65f /pretyping | |
parent | 0eff88d5a9ad9279a4e68fdb6e210c6ea671b613 (diff) |
petits changements cosmetiques sur les tactiques
+ Clear independant de l'ordre des hypotheses, et substituant les hypotheses
definies
git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@2481 85f007b7-540e-0410-9357-904b9bb8a0f7
Diffstat (limited to 'pretyping')
-rw-r--r-- | pretyping/evarutil.ml | 2 | ||||
-rw-r--r-- | pretyping/tacred.ml | 2 | ||||
-rw-r--r-- | pretyping/termops.ml | 140 | ||||
-rw-r--r-- | pretyping/termops.mli | 25 |
4 files changed, 94 insertions, 75 deletions
diff --git a/pretyping/evarutil.ml b/pretyping/evarutil.ml index a53ecf535..c556f998e 100644 --- a/pretyping/evarutil.ml +++ b/pretyping/evarutil.ml @@ -160,7 +160,7 @@ let split_evar_to_arrow sigma (ev,args) = let (sigma1,dom) = new_type_var evenv sigma in let hyps = evd.evar_hyps in let nvar = next_ident_away (id_of_string "x") (ids_of_named_context hyps) in - let newenv = push_named_decl (nvar, None, dom) evenv in + let newenv = push_named (nvar, None, dom) evenv in let (sigma2,rng) = new_type_var newenv sigma1 in let x = named_hd newenv dom Anonymous in let prod = mkProd (x, dom, subst_var nvar rng) in diff --git a/pretyping/tacred.ml b/pretyping/tacred.ml index 151525c31..ad7e02b77 100644 --- a/pretyping/tacred.ml +++ b/pretyping/tacred.ml @@ -749,7 +749,7 @@ let abstract_scheme env (locc,a,ta) t = if occur_meta a then mkLambda (na,ta,t) else - mkLambda (na, ta,subst_term_occ locc a t) + mkLambda (na, ta,subst_term_occ env locc a t) let pattern_occs loccs_trm_typ env sigma c = diff --git a/pretyping/termops.ml b/pretyping/termops.ml index 8958d1f58..867bbf50f 100644 --- a/pretyping/termops.ml +++ b/pretyping/termops.ml @@ -13,9 +13,9 @@ open Util open Names open Nameops open Term +open Sign open Environ open Nametab -open Sign let print_sort = function | Prop Pos -> (str "Set") @@ -86,7 +86,7 @@ let push_named_rec_types (lna,typarray,_) env = | Anonymous -> anomaly "Fix declarations must be named") lna typarray in Array.fold_left - (fun e assum -> push_named_decl assum e) env ctxt + (fun e assum -> push_named assum e) env ctxt let rec lookup_rel_id id sign = let rec lookrec = function @@ -335,25 +335,34 @@ let dependent m t = (* substitution functions *) (***************************) +let rec subst_meta bl c = + match kind_of_term c with + | Meta i -> (try List.assoc i bl with Not_found -> c) + | _ -> map_constr (subst_meta bl) c + (* Equality modulo let reduction *) -let rec whd_rel hyps c = +let rec whd_locals env c = match kind_of_term c with Rel i -> (try - (match Sign.lookup_rel i hyps with - (_,Some v,_) -> whd_rel hyps (lift i v) + (match lookup_rel i env with + (_,Some v,_) -> whd_locals env (lift i v) + | _ -> c) + with Not_found -> c) + | Var id -> + (try + (match lookup_named id env with + (_,Some v,_) -> whd_locals env v | _ -> c) with Not_found -> c) | _ -> c (* Expand de Bruijn indices bound to a value *) -let rec nf_rel hyps c = - map_constr_with_full_binders Sign.add_rel_decl nf_rel hyps (whd_rel hyps c) +let rec nf_locals env c = + map_constr_with_full_binders push_rel nf_locals env (whd_locals env c) -(* [m] is not evaluated because it is called only with terms for [m] which - have been lifted of the length of [hyps], hence [nf_rel] would have no - effect. *) -let compare_zeta hyps m n = zeta_eq_constr m (nf_rel hyps n) +(* compare terms modulo expansion of let and local variables *) +let compare_zeta env m n = zeta_eq_constr m (nf_locals env n) (* First utilities for avoiding telescope computation for subst_term *) @@ -388,7 +397,7 @@ let my_prefix_application eq_fun (e,k,c) (by_c : constr) (t : constr) = term [c] in a term [t] *) (*i Bizarre : si on cherche un sous terme clos, pourquoi le lifter ? i*) -let subst_term_gen eq_fun c t = +let subst_term_gen eq_fun env c t = let rec substrec (e,k,c as kc) t = match prefix_application eq_fun kc t with | Some x -> x @@ -396,43 +405,38 @@ let subst_term_gen eq_fun c t = (if eq_fun e c t then mkRel k else map_constr_with_full_binders - (fun d (e,k,c) -> (Sign.add_rel_decl d e,k+1,lift 1 c)) + (fun d (e,k,c) -> (push_rel d e,k+1,lift 1 c)) substrec kc t) in - substrec (empty_rel_context,1,c) t + substrec (env,1,nf_locals env c) t (* Recognizing occurrences of a given (closed) subterm in a term : [replace_term c1 c2 t] substitutes [c2] for all occurrences of (closed) term [c1] in a term [t] *) (*i Meme remarque : a priori [c] n'est pas forcement clos i*) -let replace_term_gen eq_fun c by_c in_t = +let replace_term_gen eq_fun env c by_c in_t = let rec substrec (e,k,c as kc) t = match my_prefix_application eq_fun kc by_c t with | Some x -> x | None -> (if eq_fun e c t then (lift k by_c) else map_constr_with_full_binders - (fun d (e,k,c) -> (Sign.add_rel_decl d e,k+1,lift 1 c)) + (fun d (e,k,c) -> (push_rel d e,k+1,lift 1 c)) substrec kc t) in - substrec (empty_rel_context,0,c) in_t + substrec (env,0,nf_locals env c) in_t -let subst_term = subst_term_gen (fun _ -> eq_constr) +let subst_term = subst_term_gen (fun _ -> eq_constr) empty_env -let replace_term = replace_term_gen (fun _ -> eq_constr) - -let rec subst_meta bl c = - match kind_of_term c with - | Meta i -> (try List.assoc i bl with Not_found -> c) - | _ -> map_constr (subst_meta bl) c +let replace_term = replace_term_gen (fun _ -> eq_constr) empty_env (* Substitute only a list of locations locs, the empty list is interpreted as substitute all, if 0 is in the list then no substitution is done. The list may contain only negative occurrences that will not be substituted. *) -let subst_term_occ_gen locs occ c t = +let subst_term_occ_gen env locs occ c t = let maxocc = List.fold_right max locs 0 in let pos = ref occ in let check = ref true in @@ -452,34 +456,34 @@ let subst_term_occ_gen locs occ c t = in incr pos; r else map_constr_with_binders_left_to_right - (fun d (e,k,c) -> (Sign.add_rel_decl d e,k+1,lift 1 c)) + (fun d (e,k,c) -> (push_rel d e,k+1,lift 1 c)) substrec kc t in - let t' = substrec (empty_rel_context,1,c) t in + let t' = substrec (env,1,nf_locals env c) t in (!pos, t') -let subst_term_occ locs c t = +let subst_term_occ env locs c t = if locs = [] then - subst_term_gen compare_zeta c t + subst_term_gen compare_zeta env c t else if List.mem 0 locs then t else - let (nbocc,t') = subst_term_occ_gen locs 1 c t in + let (nbocc,t') = subst_term_occ_gen env locs 1 c t in if List.exists (fun o -> o >= nbocc or o <= -nbocc) locs then errorlabstrm "subst_term_occ" (str "Too few occurences"); t' -let subst_term_occ_decl locs c (id,bodyopt,typ as d) = +let subst_term_occ_decl env locs c (id,bodyopt,typ as d) = match bodyopt with - | None -> (id,None,subst_term_occ locs c typ) + | None -> (id,None,subst_term_occ env locs c typ) | Some body -> if locs = [] then (id,Some (subst_term c body),type_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') = subst_term_occ_gen locs nbocc c typ in + let (nbocc,body') = subst_term_occ_gen env locs 1 c body in + let (nbocc',t') = subst_term_occ_gen env 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') @@ -685,37 +689,6 @@ let eta_eq_constr = in aux -(* Remark: Anonymous var may be dependent in Evar's contexts *) -let concrete_name env l env_names n c = - if n = Anonymous & noccurn 1 c then - (None,l) - else - let fresh_id = next_name_not_occuring env n l env_names c in - let idopt = if noccurn 1 c then None else (Some fresh_id) in - (idopt, fresh_id::l) - -let concrete_let_name env l env_names n c = - let fresh_id = next_name_not_occuring env n l env_names c in - (Name fresh_id, fresh_id::l) - -let global_vars env ids = Idset.elements (global_vars_set env ids) - -let rec rename_bound_var env l c = - match kind_of_term c with - | Prod (Name s,c1,c2) -> - if noccurn 1 c2 then - let env' = push_rel (Name s,None,c1) env in - mkProd (Name s, c1, rename_bound_var env' l c2) - else - let s' = next_ident_away s (global_vars env c2@l) in - let env' = push_rel (Name s',None,c1) env in - mkProd (Name s', c1, rename_bound_var env' (s'::l) c2) - | Prod (Anonymous,c1,c2) -> - let env' = push_rel (Anonymous,None,c1) env in - mkProd (Anonymous, c1, rename_bound_var env' l c2) - | Cast (c,t) -> mkCast (rename_bound_var env l c, t) - | x -> c - (* iterator on rel context *) let process_rel_context f env = let sign = named_context env in @@ -755,3 +728,40 @@ let make_all_name_different env = avoid := id::!avoid; push_rel (Name id,c,t) newenv) env + +let global_vars env ids = Idset.elements (global_vars_set env ids) + +let global_vars_set_of_decl env = function + | (_,None,t) -> global_vars_set env (body_of_type t) + | (_,Some c,t) -> + Idset.union (global_vars_set env (body_of_type t)) + (global_vars_set env c) + +(* Remark: Anonymous var may be dependent in Evar's contexts *) +let concrete_name env l env_names n c = + if n = Anonymous & noccurn 1 c then + (None,l) + else + let fresh_id = next_name_not_occuring env n l env_names c in + let idopt = if noccurn 1 c then None else (Some fresh_id) in + (idopt, fresh_id::l) + +let concrete_let_name env l env_names n c = + let fresh_id = next_name_not_occuring env n l env_names c in + (Name fresh_id, fresh_id::l) + +let rec rename_bound_var env l c = + match kind_of_term c with + | Prod (Name s,c1,c2) -> + if noccurn 1 c2 then + let env' = push_rel (Name s,None,c1) env in + mkProd (Name s, c1, rename_bound_var env' l c2) + else + let s' = next_ident_away s (global_vars env c2@l) in + let env' = push_rel (Name s',None,c1) env in + mkProd (Name s', c1, rename_bound_var env' (s'::l) c2) + | Prod (Anonymous,c1,c2) -> + let env' = push_rel (Anonymous,None,c1) env in + mkProd (Anonymous, c1, rename_bound_var env' l c2) + | Cast (c,t) -> mkCast (rename_bound_var env l c, t) + | x -> c diff --git a/pretyping/termops.mli b/pretyping/termops.mli index ae28029c5..1bb74a303 100644 --- a/pretyping/termops.mli +++ b/pretyping/termops.mli @@ -69,21 +69,28 @@ val occur_var_in_decl : identifier -> 'a * types option * types -> bool val free_rels : constr -> Intset.t +(* Substitution of metavariables *) +val subst_meta : (int * constr) list -> constr -> constr + +(* Expansion of local definitions *) +val whd_locals : env -> constr -> constr +val nf_locals : env -> constr -> constr + (* substitution of an arbitrary large term. Uses equality modulo reduction of let *) val dependent : constr -> constr -> bool val subst_term_gen : - (rel_context -> constr -> constr -> bool) -> constr -> constr -> constr + (env -> constr -> constr -> bool) -> env -> constr -> constr -> constr val replace_term_gen : - (rel_context -> constr -> constr -> bool) -> constr -> constr -> constr -> constr + (env -> constr -> constr -> bool) -> + env -> constr -> constr -> constr -> constr val subst_term : constr -> constr -> constr val replace_term : constr -> constr -> constr -> constr -val subst_meta : (int * constr) list -> constr -> constr val subst_term_occ_gen : - int list -> int -> constr -> types -> int * types -val subst_term_occ : int list -> constr -> types -> types + env -> int list -> int -> constr -> types -> int * types +val subst_term_occ : env -> int list -> constr -> types -> types val subst_term_occ_decl : - int list -> constr -> named_declaration -> named_declaration + env -> int list -> constr -> named_declaration -> named_declaration (* Alternative term equalities *) val zeta_eq_constr : constr -> constr -> bool @@ -136,7 +143,6 @@ val concrete_name : val concrete_let_name : env -> identifier list -> name list -> name -> constr -> name * identifier list -val global_vars : env -> constr -> identifier list val rename_bound_var : env -> identifier list -> types -> types (* other signature iterators *) @@ -144,7 +150,10 @@ val process_rel_context : (rel_declaration -> env -> env) -> env -> env val assums_of_rel_context : rel_context -> (name * constr) list val lift_rel_context : int -> rel_context -> rel_context val fold_named_context_both_sides : - ('a -> named_declaration -> named_context -> 'a) -> + ('a -> named_declaration -> named_declaration list -> 'a) -> named_context -> init:'a -> 'a val mem_named_context : identifier -> named_context -> bool val make_all_name_different : env -> env + +val global_vars : env -> constr -> identifier list +val global_vars_set_of_decl : env -> named_declaration -> Idset.t |