diff options
Diffstat (limited to 'plugins')
71 files changed, 5095 insertions, 5095 deletions
diff --git a/plugins/btauto/refl_btauto.ml b/plugins/btauto/refl_btauto.ml index 615e9cd140..ec9f9a39e0 100644 --- a/plugins/btauto/refl_btauto.ml +++ b/plugins/btauto/refl_btauto.ml @@ -54,10 +54,10 @@ module Env = struct type t = (int ConstrHashtbl.t * int ref) let add (tbl, off) (t : Constr.t) = - try ConstrHashtbl.find tbl t + try ConstrHashtbl.find tbl t with - | Not_found -> - let i = !off in + | Not_found -> + let i = !off in let () = ConstrHashtbl.add tbl t i in let () = incr off in i @@ -159,7 +159,7 @@ module Btauto = struct | Bool.Xorb (b1, b2) -> lapp f_xor [|convert b1; convert b2|] | Bool.Ifb (b1, b2, b3) -> lapp f_ifb [|convert b1; convert b2; convert b3|] - let convert_env env : Constr.t = + let convert_env env : Constr.t = CoqList.of_list (Lazy.force Bool.typ) env let reify env t = lapp eval [|convert_env env; convert t|] diff --git a/plugins/cc/ccalgo.ml b/plugins/cc/ccalgo.ml index 6f8fe8959c..500f464ea7 100644 --- a/plugins/cc/ccalgo.ml +++ b/plugins/cc/ccalgo.ml @@ -56,7 +56,7 @@ module ST=struct module IntPairTable = Hashtbl.Make(IntPair) type t = {toterm: int IntPairTable.t; - tosign: (int * int) IntTable.t} + tosign: (int * int) IntTable.t} let empty () = {toterm=IntPairTable.create init_size; @@ -64,19 +64,19 @@ module ST=struct let enter t sign st= if IntPairTable.mem st.toterm sign then - anomaly ~label:"enter" (Pp.str "signature already entered.") + anomaly ~label:"enter" (Pp.str "signature already entered.") else - IntPairTable.replace st.toterm sign t; - IntTable.replace st.tosign t sign + IntPairTable.replace st.toterm sign t; + IntTable.replace st.tosign t sign let query sign st=IntPairTable.find st.toterm sign let delete st t= try let sign=IntTable.find st.tosign t in - IntPairTable.remove st.toterm sign; - IntTable.remove st.tosign t + IntPairTable.remove st.toterm sign; + IntTable.remove st.tosign t with - Not_found -> () + Not_found -> () let delete_set st s = Int.Set.iter (delete st) s @@ -199,7 +199,7 @@ type quant_eq = qe_rhs: ccpattern; qe_rhs_valid:patt_kind } - + let swap eq : equality = let swap_rule=match eq.rule with Congruence -> Congruence @@ -234,21 +234,21 @@ type node = module Constrhash = Hashtbl.Make (struct type t = constr - let equal = eq_constr_nounivs - let hash = Constr.hash + let equal = eq_constr_nounivs + let hash = Constr.hash end) module Typehash = Constrhash module Termhash = Hashtbl.Make (struct type t = term - let equal = term_equal - let hash = hash_term + let equal = term_equal + let hash = hash_term end) module Identhash = Hashtbl.Make (struct type t = Id.t - let equal = Id.equal - let hash = Id.hash + let equal = Id.equal + let hash = Id.hash end) type forest= @@ -293,7 +293,7 @@ let empty_forest() = axioms=Constrhash.create init_size; syms=Termhash.create init_size } - + let empty depth gls:state = { uf= empty_forest (); @@ -311,7 +311,7 @@ let empty depth gls:state = env=pf_env gls; sigma=project gls } - + let forest state = state.uf let compress_path uf i j = uf.map.(j).cpath<-i @@ -332,11 +332,11 @@ let get_constructors uf i= uf.map.(i).constructors let rec find_oldest_pac uf i pac= try PacMap.find pac (get_constructors uf i) with - Not_found -> - match uf.map.(i).clas with - Eqto (j,_) -> find_oldest_pac uf j pac + Not_found -> + match uf.map.(i).clas with + Eqto (j,_) -> find_oldest_pac uf j pac | Rep _ -> raise Not_found - + let get_constructor_info uf i= match uf.map.(i).term with @@ -397,11 +397,11 @@ let next uf= if Int.equal nsize uf.max_size then let newmax=uf.max_size * 3 / 2 + 1 in let newmap=Array.make newmax dummy_node in - begin - uf.max_size<-newmax; - Array.blit uf.map 0 newmap 0 size; - uf.map<-newmap - end + begin + uf.max_size<-newmax; + Array.blit uf.map 0 newmap 0 size; + uf.map<-newmap + end else (); uf.size<-nsize; size @@ -440,14 +440,14 @@ let rec canonize_name sigma c = let func c = canonize_name sigma (EConstr.of_constr c) in match Constr.kind c with | Const (kn,u) -> - let canon_const = Constant.make1 (Constant.canonical kn) in - (mkConstU (canon_const,u)) + let canon_const = Constant.make1 (Constant.canonical kn) in + (mkConstU (canon_const,u)) | Ind ((kn,i),u) -> - let canon_mind = MutInd.make1 (MutInd.canonical kn) in - (mkIndU ((canon_mind,i),u)) + let canon_mind = MutInd.make1 (MutInd.canonical kn) in + (mkIndU ((canon_mind,i),u)) | Construct (((kn,i),j),u) -> - let canon_mind = MutInd.make1 (MutInd.canonical kn) in - mkConstructU (((canon_mind,i),j),u) + let canon_mind = MutInd.make1 (MutInd.canonical kn) in + mkConstructU (((canon_mind,i),j),u) | Prod (na,t,ct) -> mkProd (na,func t, func ct) | Lambda (na,t,ct) -> @@ -457,9 +457,9 @@ let rec canonize_name sigma c = | App (ct,l) -> mkApp (func ct,Array.Smart.map func l) | Proj(p,c) -> - let p' = Projection.map (fun kn -> + let p' = Projection.map (fun kn -> MutInd.make1 (MutInd.canonical kn)) p in - (mkProj (p', func c)) + (mkProj (p', func c)) | _ -> c (* rebuild a term from a pattern and a substitution *) @@ -477,8 +477,8 @@ let rec inst_pattern subst = function subst.(pred i) | PApp (t, args) -> List.fold_right - (fun spat f -> Appli (f,inst_pattern subst spat)) - args t + (fun spat f -> Appli (f,inst_pattern subst spat)) + args t let pr_idx_term env sigma uf i = str "[" ++ int i ++ str ":=" ++ Printer.pr_econstr_env env sigma (EConstr.of_constr (constr_of_term (term uf i))) ++ str "]" @@ -489,62 +489,62 @@ let pr_term env sigma t = str "[" ++ let rec add_term state t= let uf=state.uf in try Termhash.find uf.syms t with - Not_found -> - let b=next uf in + Not_found -> + let b=next uf in let trm = constr_of_term t in let typ = Typing.unsafe_type_of state.env state.sigma (EConstr.of_constr trm) in let typ = canonize_name state.sigma typ in - let new_node= - match t with - Symb _ | Product (_,_) -> - let paf = - {fsym=b; - fnargs=0} in - Queue.add (b,Fmark paf) state.marks; - {clas= Rep (new_representative typ); - cpath= -1; - constructors=PacMap.empty; - vertex= Leaf; - term= t} - | Eps id -> - {clas= Rep (new_representative typ); - cpath= -1; - constructors=PacMap.empty; - vertex= Leaf; - term= t} - | Appli (t1,t2) -> - let i1=add_term state t1 and i2=add_term state t2 in - add_lfather uf (find uf i1) b; - add_rfather uf (find uf i2) b; - state.terms<-Int.Set.add b state.terms; - {clas= Rep (new_representative typ); - cpath= -1; - constructors=PacMap.empty; - vertex= Node(i1,i2); - term= t} - | Constructor cinfo -> - let paf = - {fsym=b; - fnargs=0} in - Queue.add (b,Fmark paf) state.marks; - let pac = - {cnode= b; - arity= cinfo.ci_arity; - args=[]} in - Queue.add (b,Cmark pac) state.marks; - {clas=Rep (new_representative typ); - cpath= -1; - constructors=PacMap.empty; - vertex=Leaf; - term=t} - in - uf.map.(b)<-new_node; - Termhash.add uf.syms t b; - Typehash.replace state.by_type typ - (Int.Set.add b - (try Typehash.find state.by_type typ with - Not_found -> Int.Set.empty)); - b + let new_node= + match t with + Symb _ | Product (_,_) -> + let paf = + {fsym=b; + fnargs=0} in + Queue.add (b,Fmark paf) state.marks; + {clas= Rep (new_representative typ); + cpath= -1; + constructors=PacMap.empty; + vertex= Leaf; + term= t} + | Eps id -> + {clas= Rep (new_representative typ); + cpath= -1; + constructors=PacMap.empty; + vertex= Leaf; + term= t} + | Appli (t1,t2) -> + let i1=add_term state t1 and i2=add_term state t2 in + add_lfather uf (find uf i1) b; + add_rfather uf (find uf i2) b; + state.terms<-Int.Set.add b state.terms; + {clas= Rep (new_representative typ); + cpath= -1; + constructors=PacMap.empty; + vertex= Node(i1,i2); + term= t} + | Constructor cinfo -> + let paf = + {fsym=b; + fnargs=0} in + Queue.add (b,Fmark paf) state.marks; + let pac = + {cnode= b; + arity= cinfo.ci_arity; + args=[]} in + Queue.add (b,Cmark pac) state.marks; + {clas=Rep (new_representative typ); + cpath= -1; + constructors=PacMap.empty; + vertex=Leaf; + term=t} + in + uf.map.(b)<-new_node; + Termhash.add uf.syms t b; + Typehash.replace state.by_type typ + (Int.Set.add b + (try Typehash.find state.by_type typ with + Not_found -> Int.Set.empty)); + b let add_equality state c s t= let i = add_term state s in @@ -573,7 +573,7 @@ let is_redundant state id args = let prev_args = Identhash.find_all state.q_history id in List.exists (fun old_args -> - Util.Array.for_all2 (fun i j -> Int.equal i (find state.uf j)) + Util.Array.for_all2 (fun i j -> Int.equal i (find state.uf j)) norm_args old_args) prev_args with Not_found -> false @@ -591,26 +591,26 @@ let add_inst state (inst,int_subst) = let args = Array.map constr_of_term subst in let _ = Array.rev args in (* highest deBruijn index first *) let prf= mkApp(prfhead,args) in - let s = inst_pattern subst inst.qe_lhs - and t = inst_pattern subst inst.qe_rhs in - state.changed<-true; - state.rew_depth<-pred state.rew_depth; - if inst.qe_pol then - begin - debug (fun () -> - (str "Adding new equality, depth="++ int state.rew_depth) ++ fnl () ++ + let s = inst_pattern subst inst.qe_lhs + and t = inst_pattern subst inst.qe_rhs in + state.changed<-true; + state.rew_depth<-pred state.rew_depth; + if inst.qe_pol then + begin + debug (fun () -> + (str "Adding new equality, depth="++ int state.rew_depth) ++ fnl () ++ (str " [" ++ Printer.pr_econstr_env state.env state.sigma (EConstr.of_constr prf) ++ str " : " ++ pr_term state.env state.sigma s ++ str " == " ++ pr_term state.env state.sigma t ++ str "]")); - add_equality state prf s t - end - else - begin - debug (fun () -> - (str "Adding new disequality, depth="++ int state.rew_depth) ++ fnl () ++ + add_equality state prf s t + end + else + begin + debug (fun () -> + (str "Adding new disequality, depth="++ int state.rew_depth) ++ fnl () ++ (str " [" ++ Printer.pr_econstr_env state.env state.sigma (EConstr.of_constr prf) ++ str " : " ++ pr_term state.env state.sigma s ++ str " <> " ++ pr_term state.env state.sigma t ++ str "]")); - add_disequality state (Hyp prf) s t - end + add_disequality state (Hyp prf) s t + end end let link uf i j eq = (* links i -> j *) @@ -643,8 +643,8 @@ let union state i1 i2 eq= link state.uf i1 i2 eq; Constrhash.replace state.by_type r1.class_type (Int.Set.remove i1 - (try Constrhash.find state.by_type r1.class_type with - Not_found -> Int.Set.empty)); + (try Constrhash.find state.by_type r1.class_type with + Not_found -> Int.Set.empty)); let f= Int.Set.union r1.fathers r2.fathers in r2.weight<-Int.Set.cardinal f; r2.fathers<-f; @@ -652,28 +652,28 @@ let union state i1 i2 eq= ST.delete_set state.sigtable r1.fathers; state.terms<-Int.Set.union state.terms r1.fathers; PacMap.iter - (fun pac b -> Queue.add (b,Cmark pac) state.marks) - state.uf.map.(i1).constructors; + (fun pac b -> Queue.add (b,Cmark pac) state.marks) + state.uf.map.(i1).constructors; PafMap.iter - (fun paf -> Int.Set.iter - (fun b -> Queue.add (b,Fmark paf) state.marks)) - r1.functions; + (fun paf -> Int.Set.iter + (fun b -> Queue.add (b,Fmark paf) state.marks)) + r1.functions; match r1.inductive_status,r2.inductive_status with - Unknown,_ -> () - | Partial pac,Unknown -> - r2.inductive_status<-Partial pac; - state.pa_classes<-Int.Set.remove i1 state.pa_classes; - state.pa_classes<-Int.Set.add i2 state.pa_classes - | Partial _ ,(Partial _ |Partial_applied) -> - state.pa_classes<-Int.Set.remove i1 state.pa_classes - | Partial_applied,Unknown -> - r2.inductive_status<-Partial_applied - | Partial_applied,Partial _ -> - state.pa_classes<-Int.Set.remove i2 state.pa_classes; - r2.inductive_status<-Partial_applied - | Total cpl,Unknown -> r2.inductive_status<-Total cpl; - | Total (i,pac),Total _ -> Queue.add (i,Cmark pac) state.marks - | _,_ -> () + Unknown,_ -> () + | Partial pac,Unknown -> + r2.inductive_status<-Partial pac; + state.pa_classes<-Int.Set.remove i1 state.pa_classes; + state.pa_classes<-Int.Set.add i2 state.pa_classes + | Partial _ ,(Partial _ |Partial_applied) -> + state.pa_classes<-Int.Set.remove i1 state.pa_classes + | Partial_applied,Unknown -> + r2.inductive_status<-Partial_applied + | Partial_applied,Partial _ -> + state.pa_classes<-Int.Set.remove i2 state.pa_classes; + r2.inductive_status<-Partial_applied + | Total cpl,Unknown -> r2.inductive_status<-Total cpl; + | Total (i,pac),Total _ -> Queue.add (i,Cmark pac) state.marks + | _,_ -> () let merge eq state = (* merge and no-merge *) debug @@ -684,9 +684,9 @@ let merge eq state = (* merge and no-merge *) and j=find uf eq.rhs in if not (Int.equal i j) then if (size uf i)<(size uf j) then - union state i j eq + union state i j eq else - union state j i (swap eq) + union state j i (swap eq) let update t state = (* update 1 and 2 *) debug @@ -696,10 +696,10 @@ let update t state = (* update 1 and 2 *) let rep = get_representative state.uf i in begin match rep.inductive_status with - Partial _ -> - rep.inductive_status <- Partial_applied; - state.pa_classes <- Int.Set.remove i state.pa_classes - | _ -> () + Partial _ -> + rep.inductive_status <- Partial_applied; + state.pa_classes <- Int.Set.remove i state.pa_classes + | _ -> () end; PacMap.iter (fun pac _ -> Queue.add (t,Cmark (append_pac v pac)) state.marks) @@ -709,9 +709,9 @@ let update t state = (* update 1 and 2 *) rep.functions; try let s = ST.query sign state.sigtable in - Queue.add {lhs=t;rhs=s;rule=Congruence} state.combine + Queue.add {lhs=t;rhs=s;rule=Congruence} state.combine with - Not_found -> ST.enter t sign state.sigtable + Not_found -> ST.enter t sign state.sigtable let process_function_mark t rep paf state = add_paf rep paf t; @@ -720,35 +720,35 @@ let process_function_mark t rep paf state = let process_constructor_mark t i rep pac state = add_pac state.uf.map.(i) pac t; match rep.inductive_status with - Total (s,opac) -> - if not (Int.equal pac.cnode opac.cnode) then (* Conflict *) - raise (Discriminable (s,opac,t,pac)) - else (* Match *) - let cinfo = get_constructor_info state.uf pac.cnode in - let rec f n oargs args= - if n > 0 then - match (oargs,args) with - s1::q1,s2::q2-> - Queue.add - {lhs=s1;rhs=s2;rule=Injection(s,opac,t,pac,n)} - state.combine; - f (n-1) q1 q2 - | _-> anomaly ~label:"add_pacs" - (Pp.str "weird error in injection subterms merge.") - in f cinfo.ci_nhyps opac.args pac.args + Total (s,opac) -> + if not (Int.equal pac.cnode opac.cnode) then (* Conflict *) + raise (Discriminable (s,opac,t,pac)) + else (* Match *) + let cinfo = get_constructor_info state.uf pac.cnode in + let rec f n oargs args= + if n > 0 then + match (oargs,args) with + s1::q1,s2::q2-> + Queue.add + {lhs=s1;rhs=s2;rule=Injection(s,opac,t,pac,n)} + state.combine; + f (n-1) q1 q2 + | _-> anomaly ~label:"add_pacs" + (Pp.str "weird error in injection subterms merge.") + in f cinfo.ci_nhyps opac.args pac.args | Partial_applied | Partial _ -> (* add_pac state.uf.map.(i) pac t; *) - state.terms<-Int.Set.union rep.lfathers state.terms + state.terms<-Int.Set.union rep.lfathers state.terms | Unknown -> - if Int.equal pac.arity 0 then - rep.inductive_status <- Total (t,pac) - else - begin - (* add_pac state.uf.map.(i) pac t; *) - state.terms<-Int.Set.union rep.lfathers state.terms; - rep.inductive_status <- Partial pac; - state.pa_classes<- Int.Set.add i state.pa_classes - end + if Int.equal pac.arity 0 then + rep.inductive_status <- Total (t,pac) + else + begin + (* add_pac state.uf.map.(i) pac t; *) + state.terms<-Int.Set.union rep.lfathers state.terms; + rep.inductive_status <- Partial pac; + state.pa_classes<- Int.Set.add i state.pa_classes + end let process_mark t m state = debug @@ -756,7 +756,7 @@ let process_mark t m state = let i=find state.uf t in let rep=get_representative state.uf i in match m with - Fmark paf -> process_function_mark t rep paf state + Fmark paf -> process_function_mark t rep paf state | Cmark pac -> process_constructor_mark t i rep pac state type explanation = @@ -783,20 +783,20 @@ let check_disequalities state = let one_step state = try let eq = Queue.take state.combine in - merge eq state; - true + merge eq state; + true with Queue.Empty -> try - let (t,m) = Queue.take state.marks in - process_mark t m state; - true + let (t,m) = Queue.take state.marks in + process_mark t m state; + true with Queue.Empty -> - try - let t = Int.Set.choose state.terms in - state.terms<-Int.Set.remove t state.terms; - update t state; - true - with Not_found -> false + try + let t = Int.Set.choose state.terms in + state.terms<-Int.Set.remove t state.terms; + update t state; + true + with Not_found -> false let __eps__ = Id.of_string "_eps_" @@ -810,21 +810,21 @@ let new_state_var typ state = let complete_one_class state i= match (get_representative state.uf i).inductive_status with Partial pac -> - let rec app t typ n = - if n<=0 then t else + let rec app t typ n = + if n<=0 then t else let _,etyp,rest= destProd typ in let id = new_state_var (EConstr.of_constr etyp) state in app (Appli(t,Eps id)) (substl [mkVar id] rest) (n-1) in let _c = Typing.unsafe_type_of state.env state.sigma - (EConstr.of_constr (constr_of_term (term state.uf pac.cnode))) in + (EConstr.of_constr (constr_of_term (term state.uf pac.cnode))) in let _c = EConstr.Unsafe.to_constr _c in - let _args = - List.map (fun i -> constr_of_term (term state.uf i)) - pac.args in + let _args = + List.map (fun i -> constr_of_term (term state.uf i)) + pac.args in let typ = Term.prod_applist _c (List.rev _args) in - let ct = app (term state.uf i) typ pac.arity in - state.uf.epsilons <- pac :: state.uf.epsilons; - ignore (add_term state ct) + let ct = app (term state.uf i) typ pac.arity in + state.uf.epsilons <- pac :: state.uf.epsilons; + ignore (add_term state ct) | _ -> anomaly (Pp.str "wrong incomplete class.") let complete state = @@ -841,59 +841,59 @@ let make_fun_table state = Array.iteri (fun i inode -> if i < uf.size then match inode.clas with - Rep rep -> - PafMap.iter - (fun paf _ -> - let elem = - try PafMap.find paf !funtab - with Not_found -> Int.Set.empty in - funtab:= PafMap.add paf (Int.Set.add i elem) !funtab) - rep.functions - | _ -> ()) state.uf.map; + Rep rep -> + PafMap.iter + (fun paf _ -> + let elem = + try PafMap.find paf !funtab + with Not_found -> Int.Set.empty in + funtab:= PafMap.add paf (Int.Set.add i elem) !funtab) + rep.functions + | _ -> ()) state.uf.map; !funtab let do_match state res pb_stack = let mp=Stack.pop pb_stack in match mp.mp_stack with - [] -> - res:= (mp.mp_inst,mp.mp_subst) :: !res + [] -> + res:= (mp.mp_inst,mp.mp_subst) :: !res | (patt,cl)::remains -> - let uf=state.uf in - match patt with - PVar i -> - if mp.mp_subst.(pred i)<0 then - begin - mp.mp_subst.(pred i)<- cl; (* no aliasing problem here *) - Stack.push {mp with mp_stack=remains} pb_stack - end - else - if Int.equal mp.mp_subst.(pred i) cl then - Stack.push {mp with mp_stack=remains} pb_stack - else (* mismatch for non-linear variable in pattern *) () - | PApp (f,[]) -> - begin - try let j=Termhash.find uf.syms f in - if Int.equal (find uf j) cl then - Stack.push {mp with mp_stack=remains} pb_stack - with Not_found -> () - end - | PApp(f, ((last_arg::rem_args) as args)) -> - try - let j=Termhash.find uf.syms f in - let paf={fsym=j;fnargs=List.length args} in - let rep=get_representative uf cl in - let good_terms = PafMap.find paf rep.functions in - let aux i = - let (s,t) = signature state.uf i in - Stack.push - {mp with - mp_subst=Array.copy mp.mp_subst; - mp_stack= - (PApp(f,rem_args),s) :: - (last_arg,t) :: remains} pb_stack in - Int.Set.iter aux good_terms - with Not_found -> () + let uf=state.uf in + match patt with + PVar i -> + if mp.mp_subst.(pred i)<0 then + begin + mp.mp_subst.(pred i)<- cl; (* no aliasing problem here *) + Stack.push {mp with mp_stack=remains} pb_stack + end + else + if Int.equal mp.mp_subst.(pred i) cl then + Stack.push {mp with mp_stack=remains} pb_stack + else (* mismatch for non-linear variable in pattern *) () + | PApp (f,[]) -> + begin + try let j=Termhash.find uf.syms f in + if Int.equal (find uf j) cl then + Stack.push {mp with mp_stack=remains} pb_stack + with Not_found -> () + end + | PApp(f, ((last_arg::rem_args) as args)) -> + try + let j=Termhash.find uf.syms f in + let paf={fsym=j;fnargs=List.length args} in + let rep=get_representative uf cl in + let good_terms = PafMap.find paf rep.functions in + let aux i = + let (s,t) = signature state.uf i in + Stack.push + {mp with + mp_subst=Array.copy mp.mp_subst; + mp_stack= + (PApp(f,rem_args),s) :: + (last_arg,t) :: remains} pb_stack in + Int.Set.iter aux good_terms + with Not_found -> () let paf_of_patt syms = function PVar _ -> invalid_arg "paf_of_patt: pattern is trivial" @@ -908,49 +908,49 @@ let init_pb_stack state = let aux inst = begin let good_classes = - match inst.qe_lhs_valid with - Creates_variables -> Int.Set.empty - | Normal -> - begin - try - let paf= paf_of_patt syms inst.qe_lhs in - PafMap.find paf funtab - with Not_found -> Int.Set.empty - end - | Trivial typ -> - begin - try - Typehash.find state.by_type typ - with Not_found -> Int.Set.empty - end in - Int.Set.iter (fun i -> - Stack.push - {mp_subst = Array.make inst.qe_nvars (-1); - mp_inst=inst; - mp_stack=[inst.qe_lhs,i]} pb_stack) good_classes + match inst.qe_lhs_valid with + Creates_variables -> Int.Set.empty + | Normal -> + begin + try + let paf= paf_of_patt syms inst.qe_lhs in + PafMap.find paf funtab + with Not_found -> Int.Set.empty + end + | Trivial typ -> + begin + try + Typehash.find state.by_type typ + with Not_found -> Int.Set.empty + end in + Int.Set.iter (fun i -> + Stack.push + {mp_subst = Array.make inst.qe_nvars (-1); + mp_inst=inst; + mp_stack=[inst.qe_lhs,i]} pb_stack) good_classes end; begin let good_classes = - match inst.qe_rhs_valid with - Creates_variables -> Int.Set.empty - | Normal -> - begin - try - let paf= paf_of_patt syms inst.qe_rhs in - PafMap.find paf funtab - with Not_found -> Int.Set.empty - end - | Trivial typ -> - begin - try - Typehash.find state.by_type typ - with Not_found -> Int.Set.empty - end in - Int.Set.iter (fun i -> - Stack.push - {mp_subst = Array.make inst.qe_nvars (-1); - mp_inst=inst; - mp_stack=[inst.qe_rhs,i]} pb_stack) good_classes + match inst.qe_rhs_valid with + Creates_variables -> Int.Set.empty + | Normal -> + begin + try + let paf= paf_of_patt syms inst.qe_rhs in + PafMap.find paf funtab + with Not_found -> Int.Set.empty + end + | Trivial typ -> + begin + try + Typehash.find state.by_type typ + with Not_found -> Int.Set.empty + end in + Int.Set.iter (fun i -> + Stack.push + {mp_subst = Array.make inst.qe_nvars (-1); + mp_inst=inst; + mp_stack=[inst.qe_rhs,i]} pb_stack) good_classes end in List.iter aux state.quant; pb_stack @@ -962,8 +962,8 @@ let find_instances state = debug (fun () -> str "Running E-matching algorithm ... "); try while true do - Control.check_for_interrupt (); - do_match state res pb_stack + Control.check_for_interrupt (); + do_match state res pb_stack done; anomaly (Pp.str "get out of here!") with Stack.Empty -> () in @@ -977,37 +977,37 @@ let rec execute first_run state = one_step state do () done; match check_disequalities state with - None -> - if not(Int.Set.is_empty state.pa_classes) then - begin - debug (fun () -> str "First run was incomplete, completing ... "); - complete state; - execute false state - end - else - if state.rew_depth>0 then - let l=find_instances state in - List.iter (add_inst state) l; - if state.changed then - begin - state.changed <- false; - execute true state - end - else - begin - debug (fun () -> str "Out of instances ... "); - None - end - else - begin - debug (fun () -> str "Out of depth ... "); - None - end + None -> + if not(Int.Set.is_empty state.pa_classes) then + begin + debug (fun () -> str "First run was incomplete, completing ... "); + complete state; + execute false state + end + else + if state.rew_depth>0 then + let l=find_instances state in + List.iter (add_inst state) l; + if state.changed then + begin + state.changed <- false; + execute true state + end + else + begin + debug (fun () -> str "Out of instances ... "); + None + end + else + begin + debug (fun () -> str "Out of depth ... "); + None + end | Some dis -> Some - begin - if first_run then Contradiction dis - else Incomplete - end + begin + if first_run then Contradiction dis + else Incomplete + end with Discriminable(s,spac,t,tpac) -> Some begin if first_run then Discrimination (s,spac,t,tpac) diff --git a/plugins/cc/ccproof.ml b/plugins/cc/ccproof.ml index f47a14cdc7..f82a55fe71 100644 --- a/plugins/cc/ccproof.ml +++ b/plugins/cc/ccproof.ml @@ -43,25 +43,25 @@ let rec ptrans p1 p3= | Trans(p1,p2), _ ->ptrans p1 (ptrans p2 p3) | Congr(p1,p2), Congr(p3,p4) ->pcongr (ptrans p1 p3) (ptrans p2 p4) | Congr(p1,p2), Trans({p_rule=Congr(p3,p4)},p5) -> - ptrans (pcongr (ptrans p1 p3) (ptrans p2 p4)) p5 + ptrans (pcongr (ptrans p1 p3) (ptrans p2 p4)) p5 | _, _ -> if term_equal p1.p_rhs p3.p_lhs then - {p_lhs=p1.p_lhs; - p_rhs=p3.p_rhs; - p_rule=Trans (p1,p3)} + {p_lhs=p1.p_lhs; + p_rhs=p3.p_rhs; + p_rule=Trans (p1,p3)} else anomaly (Pp.str "invalid cc transitivity.") let rec psym p = match p.p_rule with Refl _ -> p | SymAx s -> - {p_lhs=p.p_rhs; - p_rhs=p.p_lhs; - p_rule=Ax s} + {p_lhs=p.p_rhs; + p_rhs=p.p_lhs; + p_rule=Ax s} | Ax s-> - {p_lhs=p.p_rhs; - p_rhs=p.p_lhs; - p_rule=SymAx s} + {p_lhs=p.p_rhs; + p_rhs=p.p_lhs; + p_rule=SymAx s} | Inject (p0,c,n,a)-> {p_lhs=p.p_rhs; p_rhs=p.p_lhs; @@ -84,9 +84,9 @@ let psymax axioms s = let rec nth_arg t n= match t with Appli (t1,t2)-> - if n>0 then - nth_arg t1 (n-1) - else t2 + if n>0 then + nth_arg t1 (n-1) + else t2 | _ -> anomaly ~label:"nth_arg" (Pp.str "not enough args.") let pinject p c n a = @@ -99,7 +99,7 @@ let rec equal_proof env sigma uf i j= if i=j then prefl (term uf i) else let (li,lj)=join_path uf i j in ptrans (path_proof env sigma uf i li) (psym (path_proof env sigma uf j lj)) - + and edge_proof env sigma uf ((i,j),eq)= debug (fun () -> str "edge_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ pr_idx_term env sigma uf j); let pi=equal_proof env sigma uf i eq.lhs in @@ -107,15 +107,15 @@ and edge_proof env sigma uf ((i,j),eq)= let pij= match eq.rule with Axiom (s,reversed)-> - if reversed then psymax (axioms uf) s - else pax (axioms uf) s + if reversed then psymax (axioms uf) s + else pax (axioms uf) s | Congruence ->congr_proof env sigma uf eq.lhs eq.rhs - | Injection (ti,ipac,tj,jpac,k) -> (* pi_k ipac = p_k jpac *) + | Injection (ti,ipac,tj,jpac,k) -> (* pi_k ipac = p_k jpac *) let p=ind_proof env sigma uf ti ipac tj jpac in let cinfo= get_constructor_info uf ipac.cnode in - pinject p cinfo.ci_constr cinfo.ci_nhyps k in + pinject p cinfo.ci_constr cinfo.ci_nhyps k in ptrans (ptrans pi pij) pj - + and constr_proof env sigma uf i ipac= debug (fun () -> str "constr_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20)); let t=find_oldest_pac uf i ipac in @@ -128,20 +128,20 @@ and constr_proof env sigma uf i ipac= let targ=term uf arg in let p=constr_proof env sigma uf fi fipac in ptrans eq_it (pcongr p (prefl targ)) - + and path_proof env sigma uf i l= debug (fun () -> str "path_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ str "{" ++ - (prlist_with_sep (fun () -> str ",") (fun ((_,j),_) -> int j) l) ++ str "}"); + (prlist_with_sep (fun () -> str ",") (fun ((_,j),_) -> int j) l) ++ str "}"); match l with | [] -> prefl (term uf i) | x::q->ptrans (path_proof env sigma uf (snd (fst x)) q) (edge_proof env sigma uf x) - + and congr_proof env sigma uf i j= debug (fun () -> str "congr_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ pr_idx_term env sigma uf j); let (i1,i2) = subterms uf i and (j1,j2) = subterms uf j in pcongr (equal_proof env sigma uf i1 j1) (equal_proof env sigma uf i2 j2) - + and ind_proof env sigma uf i ipac j jpac= debug (fun () -> str "ind_proof " ++ pr_idx_term env sigma uf i ++ brk (1,20) ++ pr_idx_term env sigma uf j); let p=equal_proof env sigma uf i j diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml index b5be1cdd89..556e6b48e6 100644 --- a/plugins/cc/cctac.ml +++ b/plugins/cc/cctac.ml @@ -51,40 +51,40 @@ let sf_of env sigma c = snd (sort_of env sigma c) let rec decompose_term env sigma t= match EConstr.kind sigma (whd env sigma t) with App (f,args)-> - let tf=decompose_term env sigma f in - let targs=Array.map (decompose_term env sigma) args in - Array.fold_left (fun s t->Appli (s,t)) tf targs + let tf=decompose_term env sigma f in + let targs=Array.map (decompose_term env sigma) args in + Array.fold_left (fun s t->Appli (s,t)) tf targs | Prod (_,a,_b) when noccurn sigma 1 _b -> - let b = Termops.pop _b in - let sort_b = sf_of env sigma b in - let sort_a = sf_of env sigma a in - Appli(Appli(Product (sort_a,sort_b) , - decompose_term env sigma a), - decompose_term env sigma b) + let b = Termops.pop _b in + let sort_b = sf_of env sigma b in + let sort_a = sf_of env sigma a in + Appli(Appli(Product (sort_a,sort_b) , + decompose_term env sigma a), + decompose_term env sigma b) | Construct c -> - let (((mind,i_ind),i_con),u)= c in - let u = EInstance.kind sigma u in - let canon_mind = MutInd.make1 (MutInd.canonical mind) in - let canon_ind = canon_mind,i_ind in - let (oib,_)=Global.lookup_inductive (canon_ind) in + let (((mind,i_ind),i_con),u)= c in + let u = EInstance.kind sigma u in + let canon_mind = MutInd.make1 (MutInd.canonical mind) in + let canon_ind = canon_mind,i_ind in + let (oib,_)=Global.lookup_inductive (canon_ind) in let nargs=constructor_nallargs env (canon_ind,i_con) in - Constructor {ci_constr= ((canon_ind,i_con),u); - ci_arity=nargs; - ci_nhyps=nargs-oib.mind_nparams} - | Ind c -> - let (mind,i_ind),u = c in - let u = EInstance.kind sigma u in - let canon_mind = MutInd.make1 (MutInd.canonical mind) in - let canon_ind = canon_mind,i_ind in (Symb (Constr.mkIndU (canon_ind,u))) - | Const (c,u) -> - let u = EInstance.kind sigma u in - let canon_const = Constant.make1 (Constant.canonical c) in - (Symb (Constr.mkConstU (canon_const,u))) - | Proj (p, c) -> + Constructor {ci_constr= ((canon_ind,i_con),u); + ci_arity=nargs; + ci_nhyps=nargs-oib.mind_nparams} + | Ind c -> + let (mind,i_ind),u = c in + let u = EInstance.kind sigma u in + let canon_mind = MutInd.make1 (MutInd.canonical mind) in + let canon_ind = canon_mind,i_ind in (Symb (Constr.mkIndU (canon_ind,u))) + | Const (c,u) -> + let u = EInstance.kind sigma u in + let canon_const = Constant.make1 (Constant.canonical c) in + (Symb (Constr.mkConstU (canon_const,u))) + | Proj (p, c) -> let canon_mind kn = MutInd.make1 (MutInd.canonical kn) in let p' = Projection.map canon_mind p in - let c = Retyping.expand_projection env sigma p' c [] in - decompose_term env sigma c + let c = Retyping.expand_projection env sigma p' c [] in + decompose_term env sigma c | _ -> let t = Termops.strip_outer_cast sigma t in if closed0 sigma t then Symb (EConstr.to_constr ~abort_on_undefined_evars:false sigma t) else raise Not_found @@ -97,33 +97,33 @@ let atom_of_constr env sigma term = let kot = EConstr.kind sigma wh in match kot with App (f,args)-> - if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3 - then `Eq (args.(0), - decompose_term env sigma args.(1), - decompose_term env sigma args.(2)) - else `Other (decompose_term env sigma term) + if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3 + then `Eq (args.(0), + decompose_term env sigma args.(1), + decompose_term env sigma args.(2)) + else `Other (decompose_term env sigma term) | _ -> `Other (decompose_term env sigma term) let rec pattern_of_constr env sigma c = match EConstr.kind sigma (whd env sigma c) with App (f,args)-> - let pf = decompose_term env sigma f in - let pargs,lrels = List.split - (Array.map_to_list (pattern_of_constr env sigma) args) in - PApp (pf,List.rev pargs), - List.fold_left Int.Set.union Int.Set.empty lrels + let pf = decompose_term env sigma f in + let pargs,lrels = List.split + (Array.map_to_list (pattern_of_constr env sigma) args) in + PApp (pf,List.rev pargs), + List.fold_left Int.Set.union Int.Set.empty lrels | Prod (_,a,_b) when noccurn sigma 1 _b -> - let b = Termops.pop _b in - let pa,sa = pattern_of_constr env sigma a in - let pb,sb = pattern_of_constr env sigma b in - let sort_b = sf_of env sigma b in - let sort_a = sf_of env sigma a in - PApp(Product (sort_a,sort_b), - [pa;pb]),(Int.Set.union sa sb) + let b = Termops.pop _b in + let pa,sa = pattern_of_constr env sigma a in + let pb,sb = pattern_of_constr env sigma b in + let sort_b = sf_of env sigma b in + let sort_a = sf_of env sigma a in + PApp(Product (sort_a,sort_b), + [pa;pb]),(Int.Set.union sa sb) | Rel i -> PVar i,Int.Set.singleton i | _ -> - let pf = decompose_term env sigma c in - PApp (pf,[]),Int.Set.empty + let pf = decompose_term env sigma c in + PApp (pf,[]),Int.Set.empty let non_trivial = function PVar _ -> false @@ -132,52 +132,52 @@ let non_trivial = function let patterns_of_constr env sigma nrels term= let f,args= try destApp sigma (whd_delta env sigma term) with DestKO -> raise Not_found in - if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3 - then - let patt1,rels1 = pattern_of_constr env sigma args.(1) - and patt2,rels2 = pattern_of_constr env sigma args.(2) in - let valid1 = - if not (Int.equal (Int.Set.cardinal rels1) nrels) then Creates_variables - else if non_trivial patt1 then Normal - else Trivial (EConstr.to_constr sigma args.(0)) - and valid2 = - if not (Int.equal (Int.Set.cardinal rels2) nrels) then Creates_variables - else if non_trivial patt2 then Normal - else Trivial (EConstr.to_constr sigma args.(0)) in - if valid1 != Creates_variables - || valid2 != Creates_variables then - nrels,valid1,patt1,valid2,patt2 - else raise Not_found - else raise Not_found + if is_global sigma (Lazy.force _eq) f && Int.equal (Array.length args) 3 + then + let patt1,rels1 = pattern_of_constr env sigma args.(1) + and patt2,rels2 = pattern_of_constr env sigma args.(2) in + let valid1 = + if not (Int.equal (Int.Set.cardinal rels1) nrels) then Creates_variables + else if non_trivial patt1 then Normal + else Trivial (EConstr.to_constr sigma args.(0)) + and valid2 = + if not (Int.equal (Int.Set.cardinal rels2) nrels) then Creates_variables + else if non_trivial patt2 then Normal + else Trivial (EConstr.to_constr sigma args.(0)) in + if valid1 != Creates_variables + || valid2 != Creates_variables then + nrels,valid1,patt1,valid2,patt2 + else raise Not_found + else raise Not_found let rec quantified_atom_of_constr env sigma nrels term = match EConstr.kind sigma (whd_delta env sigma term) with Prod (id,atom,ff) -> - if is_global sigma (Lazy.force _False) ff then - let patts=patterns_of_constr env sigma nrels atom in - `Nrule patts - else + if is_global sigma (Lazy.force _False) ff then + let patts=patterns_of_constr env sigma nrels atom in + `Nrule patts + else quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma (succ nrels) ff - | _ -> - let patts=patterns_of_constr env sigma nrels term in - `Rule patts + | _ -> + let patts=patterns_of_constr env sigma nrels term in + `Rule patts let litteral_of_constr env sigma term= match EConstr.kind sigma (whd_delta env sigma term) with | Prod (id,atom,ff) -> - if is_global sigma (Lazy.force _False) ff then - match (atom_of_constr env sigma atom) with - `Eq(t,a,b) -> `Neq(t,a,b) - | `Other(p) -> `Nother(p) - else - begin - try + if is_global sigma (Lazy.force _False) ff then + match (atom_of_constr env sigma atom) with + `Eq(t,a,b) -> `Neq(t,a,b) + | `Other(p) -> `Nother(p) + else + begin + try quantified_atom_of_constr (EConstr.push_rel (RelDecl.LocalAssum (id,atom)) env) sigma 1 ff - with Not_found -> - `Other (decompose_term env sigma term) - end + with Not_found -> + `Other (decompose_term env sigma term) + end | _ -> - atom_of_constr env sigma term + atom_of_constr env sigma term (* store all equalities from the context *) @@ -191,38 +191,38 @@ let make_prb gls depth additionnal_terms = let neg_hyps =ref [] in List.iter (fun c -> - let t = decompose_term env sigma c in - ignore (add_term state t)) additionnal_terms; + let t = decompose_term env sigma c in + ignore (add_term state t)) additionnal_terms; List.iter (fun decl -> let id = NamedDecl.get_id decl in - begin - let cid=Constr.mkVar id in - match litteral_of_constr env sigma (NamedDecl.get_type decl) with - `Eq (t,a,b) -> add_equality state cid a b - | `Neq (t,a,b) -> add_disequality state (Hyp cid) a b - | `Other ph -> - List.iter - (fun (cidn,nh) -> - add_disequality state (HeqnH (cid,cidn)) ph nh) - !neg_hyps; - pos_hyps:=(cid,ph):: !pos_hyps - | `Nother nh -> - List.iter - (fun (cidp,ph) -> - add_disequality state (HeqnH (cidp,cid)) ph nh) - !pos_hyps; - neg_hyps:=(cid,nh):: !neg_hyps - | `Rule patts -> add_quant state id true patts - | `Nrule patts -> add_quant state id false patts - end) (Proofview.Goal.hyps gls); + begin + let cid=Constr.mkVar id in + match litteral_of_constr env sigma (NamedDecl.get_type decl) with + `Eq (t,a,b) -> add_equality state cid a b + | `Neq (t,a,b) -> add_disequality state (Hyp cid) a b + | `Other ph -> + List.iter + (fun (cidn,nh) -> + add_disequality state (HeqnH (cid,cidn)) ph nh) + !neg_hyps; + pos_hyps:=(cid,ph):: !pos_hyps + | `Nother nh -> + List.iter + (fun (cidp,ph) -> + add_disequality state (HeqnH (cidp,cid)) ph nh) + !pos_hyps; + neg_hyps:=(cid,nh):: !neg_hyps + | `Rule patts -> add_quant state id true patts + | `Nrule patts -> add_quant state id false patts + end) (Proofview.Goal.hyps gls); begin match atom_of_constr env sigma (pf_concl gls) with - `Eq (t,a,b) -> add_disequality state Goal a b - | `Other g -> - List.iter - (fun (idp,ph) -> - add_disequality state (HeqG idp) ph g) !pos_hyps + `Eq (t,a,b) -> add_disequality state Goal a b + | `Other g -> + List.iter + (fun (idp,ph) -> + add_disequality state (HeqG idp) ph g) !pos_hyps end; state @@ -275,7 +275,7 @@ let assert_before n c = let refresh_type env evm ty = Evarsolve.refresh_universes ~status:Evd.univ_flexible ~refreshset:true - (Some false) env evm ty + (Some false) env evm ty let refresh_universes ty k = Proofview.Goal.enter begin fun gl -> @@ -295,60 +295,60 @@ let rec proof_tac p : unit Proofview.tactic = Ax c -> exact_check (EConstr.of_constr c) | SymAx c -> let c = EConstr.of_constr c in - let l=constr_of_term p.p_lhs and - r=constr_of_term p.p_rhs in - refresh_universes (type_of l) (fun typ -> + let l=constr_of_term p.p_lhs and + r=constr_of_term p.p_rhs in + refresh_universes (type_of l) (fun typ -> app_global _sym_eq [|typ;r;l;c|] exact_check) | Refl t -> - let lr = constr_of_term t in - refresh_universes (type_of lr) (fun typ -> + let lr = constr_of_term t in + refresh_universes (type_of lr) (fun typ -> app_global _refl_equal [|typ;constr_of_term t|] exact_check) | Trans (p1,p2)-> - let t1 = constr_of_term p1.p_lhs and - t2 = constr_of_term p1.p_rhs and - t3 = constr_of_term p2.p_rhs in - refresh_universes (type_of t2) (fun typ -> - let prf = app_global_with_holes _trans_eq [|typ;t1;t2;t3;|] 2 in + let t1 = constr_of_term p1.p_lhs and + t2 = constr_of_term p1.p_rhs and + t3 = constr_of_term p2.p_rhs in + refresh_universes (type_of t2) (fun typ -> + let prf = app_global_with_holes _trans_eq [|typ;t1;t2;t3;|] 2 in Tacticals.New.tclTHENS prf [(proof_tac p1);(proof_tac p2)]) | Congr (p1,p2)-> - let tf1=constr_of_term p1.p_lhs - and tx1=constr_of_term p2.p_lhs - and tf2=constr_of_term p1.p_rhs - and tx2=constr_of_term p2.p_rhs in - refresh_universes (type_of tf1) (fun typf -> - refresh_universes (type_of tx1) (fun typx -> - refresh_universes (type_of (mkApp (tf1,[|tx1|]))) (fun typfx -> + let tf1=constr_of_term p1.p_lhs + and tx1=constr_of_term p2.p_lhs + and tf2=constr_of_term p1.p_rhs + and tx2=constr_of_term p2.p_rhs in + refresh_universes (type_of tf1) (fun typf -> + refresh_universes (type_of tx1) (fun typx -> + refresh_universes (type_of (mkApp (tf1,[|tx1|]))) (fun typfx -> let id = Tacmach.New.pf_get_new_id (Id.of_string "f") gl in let appx1 = mkLambda(make_annot (Name id) Sorts.Relevant,typf,mkApp(mkRel 1,[|tx1|])) in - let lemma1 = app_global_with_holes _f_equal [|typf;typfx;appx1;tf1;tf2|] 1 in - let lemma2 = app_global_with_holes _f_equal [|typx;typfx;tf2;tx1;tx2|] 1 in - let prf = - app_global_with_holes _trans_eq - [|typfx; - mkApp(tf1,[|tx1|]); - mkApp(tf2,[|tx1|]); - mkApp(tf2,[|tx2|])|] 2 in - Tacticals.New.tclTHENS prf - [Tacticals.New.tclTHEN lemma1 (proof_tac p1); - Tacticals.New.tclFIRST - [Tacticals.New.tclTHEN lemma2 (proof_tac p2); - reflexivity; + let lemma1 = app_global_with_holes _f_equal [|typf;typfx;appx1;tf1;tf2|] 1 in + let lemma2 = app_global_with_holes _f_equal [|typx;typfx;tf2;tx1;tx2|] 1 in + let prf = + app_global_with_holes _trans_eq + [|typfx; + mkApp(tf1,[|tx1|]); + mkApp(tf2,[|tx1|]); + mkApp(tf2,[|tx2|])|] 2 in + Tacticals.New.tclTHENS prf + [Tacticals.New.tclTHEN lemma1 (proof_tac p1); + Tacticals.New.tclFIRST + [Tacticals.New.tclTHEN lemma2 (proof_tac p2); + reflexivity; Tacticals.New.tclZEROMSG - (Pp.str - "I don't know how to handle dependent equality")]]))) + (Pp.str + "I don't know how to handle dependent equality")]]))) | Inject (prf,cstr,nargs,argind) -> - let ti=constr_of_term prf.p_lhs in - let tj=constr_of_term prf.p_rhs in - let default=constr_of_term p.p_lhs in - let special=mkRel (1+nargs-argind) in - refresh_universes (type_of ti) (fun intype -> + let ti=constr_of_term prf.p_lhs in + let tj=constr_of_term prf.p_rhs in + let default=constr_of_term p.p_lhs in + let special=mkRel (1+nargs-argind) in + refresh_universes (type_of ti) (fun intype -> refresh_universes (type_of default) (fun outtype -> let sigma, proj = build_projection intype cstr special default gl in - let injt= + let injt= app_global_with_holes _f_equal [|intype;outtype;proj;ti;tj|] 1 in - Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) + Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Tacticals.New.tclTHEN injt (proof_tac prf)))) with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e end @@ -371,7 +371,7 @@ let refine_exact_check c = Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evm) (exact_check c) end -let convert_to_goal_tac c t1 t2 p = +let convert_to_goal_tac c t1 t2 p = Proofview.Goal.enter begin fun gl -> let tt1=constr_of_term t1 and tt2=constr_of_term t2 in let k sort = @@ -381,7 +381,7 @@ let convert_to_goal_tac c t1 t2 p = let identity=mkLambda (make_annot (Name x) Sorts.Relevant,sort,mkRel 1) in let endt = app_global _eq_rect [|sort;tt1;identity;c;tt2;mkVar e|] in Tacticals.New.tclTHENS (neweq (assert_before (Name e))) - [proof_tac p; endt refine_exact_check] + [proof_tac p; endt refine_exact_check] in refresh_universes (Tacmach.New.pf_unsafe_type_of gl tt2) k end @@ -405,7 +405,7 @@ let discriminate_tac cstru p = let hid = Tacmach.New.pf_get_new_id (Id.of_string "Heq") gl in let neweq=app_global _eq [|intype;lhs;rhs|] in Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS evm) - (Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) + (Tacticals.New.tclTHENS (neweq (assert_before (Name hid))) [proof_tac p; Equality.discrHyp hid]) end @@ -430,13 +430,13 @@ let cc_tactic depth additionnal_terms = match sol with None -> Tacticals.New.tclFAIL 0 (str "congruence failed") | Some reason -> - debug (fun () -> Pp.str "Goal solved, generating proof ..."); - match reason with - Discrimination (i,ipac,j,jpac) -> + debug (fun () -> Pp.str "Goal solved, generating proof ..."); + match reason with + Discrimination (i,ipac,j,jpac) -> let p=build_proof (Tacmach.New.pf_env gl) sigma uf (`Discr (i,ipac,j,jpac)) in - let cstr=(get_constructor_info uf ipac.cnode).ci_constr in - discriminate_tac cstr p - | Incomplete -> + let cstr=(get_constructor_info uf ipac.cnode).ci_constr in + discriminate_tac cstr p + | Incomplete -> let open Glob_term in let env = Proofview.Goal.env gl in let terms_to_complete = List.map (build_term_to_complete uf) (epsilons uf) in @@ -456,20 +456,20 @@ let cc_tactic depth additionnal_terms = end ++ str " replacing metavariables by arbitrary terms.") in Tacticals.New.tclFAIL 0 msg - | Contradiction dis -> + | Contradiction dis -> let env = Proofview.Goal.env gl in let p=build_proof env sigma uf (`Prove (dis.lhs,dis.rhs)) in - let ta=term uf dis.lhs and tb=term uf dis.rhs in - match dis.rule with - Goal -> proof_tac p - | Hyp id -> refute_tac (EConstr.of_constr id) ta tb p - | HeqG id -> + let ta=term uf dis.lhs and tb=term uf dis.rhs in + match dis.rule with + Goal -> proof_tac p + | Hyp id -> refute_tac (EConstr.of_constr id) ta tb p + | HeqG id -> let id = EConstr.of_constr id in - convert_to_goal_tac id ta tb p - | HeqnH (ida,idb) -> + convert_to_goal_tac id ta tb p + | HeqnH (ida,idb) -> let ida = EConstr.of_constr ida in let idb = EConstr.of_constr idb in - convert_to_hyp_tac ida ta idb tb p + convert_to_hyp_tac ida ta idb tb p end let cc_fail = @@ -509,21 +509,21 @@ let f_equal = let cut_eq c1 c2 = try (* type_of can raise an exception *) Tacticals.New.tclTHENS - (mk_eq _eq c1 c2 Tactics.cut) - [Proofview.tclUNIT ();Tacticals.New.tclTRY ((app_global _refl_equal [||]) apply)] + (mk_eq _eq c1 c2 Tactics.cut) + [Proofview.tclUNIT ();Tacticals.New.tclTRY ((app_global _refl_equal [||]) apply)] with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e in Proofview.tclORELSE begin match EConstr.kind sigma concl with | App (r,[|_;t;t'|]) when is_global sigma (Lazy.force _eq) r -> - begin match EConstr.kind sigma t, EConstr.kind sigma t' with - | App (f,v), App (f',v') when Int.equal (Array.length v) (Array.length v') -> - let rec cuts i = - if i < 0 then Tacticals.New.tclTRY (congruence_tac 1000 []) - else Tacticals.New.tclTHENFIRST (cut_eq v.(i) v'.(i)) (cuts (i-1)) - in cuts (Array.length v - 1) - | _ -> Proofview.tclUNIT () - end + begin match EConstr.kind sigma t, EConstr.kind sigma t' with + | App (f,v), App (f',v') when Int.equal (Array.length v) (Array.length v') -> + let rec cuts i = + if i < 0 then Tacticals.New.tclTRY (congruence_tac 1000 []) + else Tacticals.New.tclTHENFIRST (cut_eq v.(i) v'.(i)) (cuts (i-1)) + in cuts (Array.length v - 1) + | _ -> Proofview.tclUNIT () + end | _ -> Proofview.tclUNIT () end begin function (e, info) -> match e with diff --git a/plugins/extraction/common.ml b/plugins/extraction/common.ml index 1c325a8d3a..2f3f42c5f6 100644 --- a/plugins/extraction/common.ml +++ b/plugins/extraction/common.ml @@ -164,9 +164,9 @@ let rename_tvars avoid l = let rec rename avoid = function | [] -> [],avoid | id :: idl -> - let id = rename_id (lowercase_id id) avoid in - let idl, avoid = rename (Id.Set.add id avoid) idl in - (id :: idl, avoid) in + let id = rename_id (lowercase_id id) avoid in + let idl, avoid = rename (Id.Set.add id avoid) idl in + (id :: idl, avoid) in fst (rename avoid l) let push_vars ids (db,avoid) = @@ -271,8 +271,8 @@ let params_ren_add, params_ren_mem = *) type visible_layer = { mp : ModPath.t; - params : ModPath.t list; - mutable content : Label.t KMap.t; } + params : ModPath.t list; + mutable content : Label.t KMap.t; } let pop_visible, push_visible, get_visible = let vis = ref [] in @@ -281,10 +281,10 @@ let pop_visible, push_visible, get_visible = match !vis with | [] -> assert false | v :: vl -> - vis := vl; - (* we save the 1st-level-content of MPfile for later use *) - if get_phase () == Impl && modular () && is_modfile v.mp - then add_mpfiles_content v.mp v.content + vis := vl; + (* we save the 1st-level-content of MPfile for later use *) + if get_phase () == Impl && modular () && is_modfile v.mp + then add_mpfiles_content v.mp v.content and push mp mps = vis := { mp = mp; params = mps; content = KMap.empty } :: !vis and get () = !vis @@ -356,9 +356,9 @@ let modfstlev_rename = with Not_found -> let s = ascii_of_id id in if is_lower s || begins_with_CoqXX s then - (add_index id 1; "Coq_"^s) + (add_index id 1; "Coq_"^s) else - (add_index id 0; s) + (add_index id 0; s) (*s Creating renaming for a [module_path] : first, the real function ... *) @@ -448,13 +448,13 @@ let visible_clash mp0 ks = | [] -> false | v :: _ when ModPath.equal v.mp mp0 -> false | v :: vis -> - let b = KMap.mem ks v.content in - if b && not (is_mp_bound mp0) then true - else begin - if b then params_ren_add mp0; - if params_lookup mp0 ks v.params then false - else clash vis - end + let b = KMap.mem ks v.content in + if b && not (is_mp_bound mp0) then true + else begin + if b then params_ren_add mp0; + if params_lookup mp0 ks v.params then false + else clash vis + end in clash (get_visible ()) (* Same, but with verbose output (and mp0 shouldn't be a MPbound) *) @@ -464,10 +464,10 @@ let visible_clash_dbg mp0 ks = | [] -> None | v :: _ when ModPath.equal v.mp mp0 -> None | v :: vis -> - try Some (v.mp,KMap.find ks v.content) - with Not_found -> - if params_lookup mp0 ks v.params then None - else clash vis + try Some (v.mp,KMap.find ks v.content) + with Not_found -> + if params_lookup mp0 ks v.params then None + else clash vis in clash (get_visible ()) (* After the 1st pass, we can decide which modules will be opened initially *) @@ -483,9 +483,9 @@ let opened_libraries () = after such an open, there's no unambiguous way to refer to objects of B. *) let to_open = List.filter - (fun mp -> - not (List.exists (fun k -> KMap.mem k (get_mpfiles_content mp)) used_ks)) - used_files + (fun mp -> + not (List.exists (fun k -> KMap.mem k (get_mpfiles_content mp)) used_ks)) + used_files in mpfiles_clear (); List.iter mpfiles_add to_open; @@ -549,18 +549,18 @@ let pp_ocaml_extern k base rls = match rls with | [] -> assert false | base_s :: rls' -> if (not (modular ())) (* Pseudo qualification with "" *) - || (List.is_empty rls') (* Case of a file A.v used as a module later *) - || (not (mpfiles_mem base)) (* Module not opened *) - || (mpfiles_clash base (fstlev_ks k rls')) (* Conflict in opened files *) - || (visible_clash base (fstlev_ks k rls')) (* Local conflict *) + || (List.is_empty rls') (* Case of a file A.v used as a module later *) + || (not (mpfiles_mem base)) (* Module not opened *) + || (mpfiles_clash base (fstlev_ks k rls')) (* Conflict in opened files *) + || (visible_clash base (fstlev_ks k rls')) (* Local conflict *) then - (* We need to fully qualify. Last clash situation is unsupported *) - match visible_clash_dbg base (Mod,base_s) with - | None -> dottify rls - | Some (mp,l) -> error_module_clash base (MPdot (mp,l)) + (* We need to fully qualify. Last clash situation is unsupported *) + match visible_clash_dbg base (Mod,base_s) with + | None -> dottify rls + | Some (mp,l) -> error_module_clash base (MPdot (mp,l)) else - (* Standard situation : object in an opened file *) - dottify rls' + (* Standard situation : object in an opened file *) + dottify rls' (* [pp_ocaml_gen] : choosing between [pp_ocaml_local] or [pp_ocaml_extern] *) @@ -568,9 +568,9 @@ let pp_ocaml_gen k mp rls olab = match common_prefix_from_list mp (get_visible_mps ()) with | Some prefix -> pp_ocaml_local k prefix mp rls olab | None -> - let base = base_mp mp in - if is_mp_bound base then pp_ocaml_bound base rls - else pp_ocaml_extern k base rls + let base = base_mp mp in + if is_mp_bound base then pp_ocaml_bound base rls + else pp_ocaml_extern k base rls (* For Haskell, things are simpler: we have removed (almost) all structures *) @@ -607,9 +607,9 @@ let pp_module mp = match mp with | MPdot (mp0,l) when ModPath.equal mp0 (top_visible_mp ()) -> (* simplest situation: definition of mp (or use in the same context) *) - (* we update the visible environment *) - let s = List.hd ls in - add_visible (Mod,s) l; s + (* we update the visible environment *) + let s = List.hd ls in + add_visible (Mod,s) l; s | _ -> pp_ocaml_gen Mod mp (List.rev ls) None (** Special hack for constants of type Ascii.ascii : if an diff --git a/plugins/extraction/extract_env.ml b/plugins/extraction/extract_env.ml index 551dbdc6fb..35110552ab 100644 --- a/plugins/extraction/extract_env.ml +++ b/plugins/extraction/extract_env.ml @@ -147,8 +147,8 @@ let check_fix env sg cb i = | Def lbody -> (match EConstr.kind sg (get_body lbody) with | Fix ((_,j),recd) when Int.equal i j -> check_arity env cb; (true,recd) - | CoFix (j,recd) when Int.equal i j -> check_arity env cb; (false,recd) - | _ -> raise Impossible) + | CoFix (j,recd) when Int.equal i j -> check_arity env cb; (false,recd) + | _ -> raise Impossible) | Undef _ | OpaqueDef _ | Primitive _ -> raise Impossible let prec_declaration_equal sg (na1, ca1, ta1) (na2, ca2, ta2) = @@ -166,14 +166,14 @@ let factor_fix env sg l cb msb = let labels = Array.make n l in List.iteri (fun j -> - function - | (l,SFBconst cb') -> + function + | (l,SFBconst cb') -> let check' = check_fix env sg cb' (j+1) in if not ((fst check : bool) == (fst check') && prec_declaration_equal sg (snd check) (snd check')) - then raise Impossible; - labels.(j+1) <- l; - | _ -> raise Impossible) msb'; + then raise Impossible; + labels.(j+1) <- l; + | _ -> raise Impossible) msb'; labels, recd, msb'' end @@ -256,8 +256,8 @@ and extract_mexpr_spec env mp1 (me_struct_o,me_alg) = match me_alg with let sg = Evd.from_env env in (match extract_with_type env' sg (EConstr.of_constr c) with (* cb may contain some kn *) - | None -> mt - | Some (vl,typ) -> + | None -> mt + | Some (vl,typ) -> type_iter_references Visit.add_ref typ; MTwith(mt,ML_With_type(idl,vl,typ))) | MEwith(me',WithMod(idl,mp))-> @@ -271,8 +271,8 @@ and extract_mexpr_spec env mp1 (me_struct_o,me_alg) = match me_alg with and extract_mexpression_spec env mp1 (me_struct,me_alg) = match me_alg with | MoreFunctor (mbid, mtb, me_alg') -> let me_struct' = match me_struct with - | MoreFunctor (mbid',_,me') when MBId.equal mbid' mbid -> me' - | _ -> assert false + | MoreFunctor (mbid',_,me') when MBId.equal mbid' mbid -> me' + | _ -> assert false in let mp = MPbound mbid in let env' = Modops.add_module_type mp mtb env in @@ -288,7 +288,7 @@ and extract_msignature_spec env mp1 reso = function let mp = MPbound mbid in let env' = Modops.add_module_type mp mtb env in MTfunsig (mbid, extract_mbody_spec env mp mtb, - extract_msignature_spec env' mp1 reso me) + extract_msignature_spec env' mp1 reso me) and extract_mbody_spec : 'a. _ -> _ -> 'a generic_module_body -> _ = fun env mp mb -> match mb.mod_type_alg with @@ -308,38 +308,38 @@ let rec extract_structure env mp reso ~all = function (try let sg = Evd.from_env env in let vl,recd,struc = factor_fix env sg l cb struc in - let vc = Array.map (make_cst reso mp) vl in - let ms = extract_structure env mp reso ~all struc in - let b = Array.exists Visit.needed_cst vc in - if all || b then + let vc = Array.map (make_cst reso mp) vl in + let ms = extract_structure env mp reso ~all struc in + let b = Array.exists Visit.needed_cst vc in + if all || b then let d = extract_fixpoint env sg vc recd in - if (not b) && (logical_decl d) then ms - else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end - else ms + if (not b) && (logical_decl d) then ms + else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end + else ms with Impossible -> - let ms = extract_structure env mp reso ~all struc in - let c = make_cst reso mp l in - let b = Visit.needed_cst c in - if all || b then - let d = extract_constant env c cb in - if (not b) && (logical_decl d) then ms - else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end - else ms) + let ms = extract_structure env mp reso ~all struc in + let c = make_cst reso mp l in + let b = Visit.needed_cst c in + if all || b then + let d = extract_constant env c cb in + if (not b) && (logical_decl d) then ms + else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end + else ms) | (l,SFBmind mib) :: struc -> let ms = extract_structure env mp reso ~all struc in let mind = make_mind reso mp l in let b = Visit.needed_ind mind in if all || b then - let d = Dind (mind, extract_inductive env mind) in - if (not b) && (logical_decl d) then ms - else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end + let d = Dind (mind, extract_inductive env mind) in + if (not b) && (logical_decl d) then ms + else begin Visit.add_decl_deps d; (l,SEdecl d) :: ms end else ms | (l,SFBmodule mb) :: struc -> let ms = extract_structure env mp reso ~all struc in let mp = MPdot (mp,l) in let all' = all || Visit.needed_mp_all mp in if all' || Visit.needed_mp mp then - (l,SEmodule (extract_module env mp ~all:all' mb)) :: ms + (l,SEmodule (extract_module env mp ~all:all' mb)) :: ms else ms | (l,SFBmodtype mtb) :: struc -> let ms = extract_structure env mp reso ~all struc in @@ -363,7 +363,7 @@ and extract_mexpr env mp = function Visit.add_mp_all mp; Miniml.MEident mp | MEapply (me, arg) -> Miniml.MEapply (extract_mexpr env mp me, - extract_mexpr env mp (MEident arg)) + extract_mexpr env mp (MEident arg)) and extract_mexpression env mp = function | NoFunctor me -> extract_mexpr env mp me @@ -373,7 +373,7 @@ and extract_mexpression env mp = function Miniml.MEfunctor (mbid, extract_mbody_spec env mp1 mtb, - extract_mexpression env' mp me) + extract_mexpression env' mp me) and extract_msignature env mp reso ~all = function | NoFunctor struc -> @@ -385,7 +385,7 @@ and extract_msignature env mp reso ~all = function Miniml.MEfunctor (mbid, extract_mbody_spec env mp1 mtb, - extract_msignature env' mp reso ~all me) + extract_msignature env' mp reso ~all me) and extract_module env mp ~all mb = (* A module has an empty [mod_expr] when : @@ -447,19 +447,19 @@ let mono_filename f = match f with | None -> None, None, default_id | Some f -> - let f = - if Filename.check_suffix f d.file_suffix then - Filename.chop_suffix f d.file_suffix - else f - in - let id = - if lang () != Haskell then default_id - else + let f = + if Filename.check_suffix f d.file_suffix then + Filename.chop_suffix f d.file_suffix + else f + in + let id = + if lang () != Haskell then default_id + else try Id.of_string (Filename.basename f) - with UserError _ -> + with UserError _ -> user_err Pp.(str "Extraction: provided filename is not a valid identifier") - in - Some (f^d.file_suffix), Option.map ((^) f) d.sig_suffix, id + in + Some (f^d.file_suffix), Option.map ((^) f) d.sig_suffix, id (* Builds a suitable filename from a module id *) @@ -494,8 +494,8 @@ let formatter dry file = if dry then Format.make_formatter (fun _ _ _ -> ()) (fun _ -> ()) else match file with - | Some f -> Topfmt.with_output_to f - | None -> Format.formatter_of_buffer buf + | Some f -> Topfmt.with_output_to f + | None -> Format.formatter_of_buffer buf in (* XXX: Fixme, this shouldn't depend on Topfmt *) (* We never want to see ellipsis ... in extracted code *) @@ -554,14 +554,14 @@ let print_structure_to_file (fn,si,mo) dry struc = let cout = open_out si in let ft = formatter false (Some cout) in begin try - set_phase Intf; - pp_with ft (d.sig_preamble mo comment opened unsafe_needs); - pp_with ft (d.pp_sig (signature_of_structure struc)); + set_phase Intf; + pp_with ft (d.sig_preamble mo comment opened unsafe_needs); + pp_with ft (d.pp_sig (signature_of_structure struc)); Format.pp_print_flush ft (); - close_out cout; + close_out cout; with reraise -> Format.pp_print_flush ft (); - close_out cout; raise reraise + close_out cout; raise reraise end; info_file si) (if dry then None else si); @@ -606,9 +606,9 @@ let rec locate_ref = function in match mpo, ro with | None, None -> Nametab.error_global_not_found qid - | None, Some r -> let refs,mps = locate_ref l in r::refs,mps - | Some mp, None -> let refs,mps = locate_ref l in refs,mp::mps - | Some mp, Some r -> + | None, Some r -> let refs,mps = locate_ref l in r::refs,mps + | Some mp, None -> let refs,mps = locate_ref l in refs,mp::mps + | Some mp, Some r -> warning_ambiguous_name (qid,mp,r); let refs,mps = locate_ref l in refs,mp::mps @@ -637,7 +637,7 @@ let separate_extraction lr = warns (); let print = function | (MPfile dir as mp, sel) as e -> - print_structure_to_file (module_filename mp) false [e] + print_structure_to_file (module_filename mp) false [e] | _ -> assert false in List.iter print struc; @@ -686,8 +686,8 @@ let extraction_library is_rec m = warns (); let print = function | (MPfile dir as mp, sel) as e -> - let dry = not is_rec && not (DirPath.equal dir dir_m) in - print_structure_to_file (module_filename mp) dry [e] + let dry = not is_rec && not (DirPath.equal dir dir_m) in + print_structure_to_file (module_filename mp) dry [e] | _ -> assert false in List.iter print struc; diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 04f5b66241..a4469b7ec1 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -243,8 +243,8 @@ let parse_ind_args si args relmax = | Kill _ :: s -> parse (i+1) j s | Keep :: s -> (match Constr.kind args.(i-1) with - | Rel k -> Int.Map.add (relmax+1-k) j (parse (i+1) (j+1) s) - | _ -> parse (i+1) (j+1) s) + | Rel k -> Int.Map.add (relmax+1-k) j (parse (i+1) (j+1) s) + | _ -> parse (i+1) (j+1) s) in parse 1 1 si (*S Extraction of a type. *) @@ -265,31 +265,31 @@ let rec extract_type env sg db j c args = extract_type env sg db j d (Array.to_list args' @ args) | Lambda (_,_,d) -> (match args with - | [] -> assert false (* A lambda cannot be a type. *) + | [] -> assert false (* A lambda cannot be a type. *) | a :: args -> extract_type env sg db j (EConstr.Vars.subst1 a d) args) | Prod (n,t,d) -> assert (List.is_empty args); let env' = push_rel_assum (n,t) env in (match flag_of_type env sg t with | (Info, Default) -> - (* Standard case: two [extract_type] ... *) + (* Standard case: two [extract_type] ... *) let mld = extract_type env' sg (0::db) j d [] in - (match expand env mld with - | Tdummy d -> Tdummy d + (match expand env mld with + | Tdummy d -> Tdummy d | _ -> Tarr (extract_type env sg db 0 t [], mld)) - | (Info, TypeScheme) when j > 0 -> - (* A new type var. *) + | (Info, TypeScheme) when j > 0 -> + (* A new type var. *) let mld = extract_type env' sg (j::db) (j+1) d [] in - (match expand env mld with - | Tdummy d -> Tdummy d - | _ -> Tarr (Tdummy Ktype, mld)) - | _,lvl -> + (match expand env mld with + | Tdummy d -> Tdummy d + | _ -> Tarr (Tdummy Ktype, mld)) + | _,lvl -> let mld = extract_type env' sg (0::db) j d [] in - (match expand env mld with - | Tdummy d -> Tdummy d - | _ -> - let reason = if lvl == TypeScheme then Ktype else Kprop in - Tarr (Tdummy reason, mld))) + (match expand env mld with + | Tdummy d -> Tdummy d + | _ -> + let reason = if lvl == TypeScheme then Ktype else Kprop in + Tarr (Tdummy reason, mld))) | Sort _ -> Tdummy Ktype (* The two logical cases. *) | _ when sort_of env sg (applistc c args) == InProp -> Tdummy Kprop | Rel n -> @@ -297,16 +297,16 @@ let rec extract_type env sg db j c args = | LocalDef (_,t,_) -> extract_type env sg db j (EConstr.Vars.lift n t) args | _ -> - (* Asks [db] a translation for [n]. *) - if n > List.length db then Tunknown - else let n' = List.nth db (n-1) in - if Int.equal n' 0 then Tunknown else Tvar n') + (* Asks [db] a translation for [n]. *) + if n > List.length db then Tunknown + else let n' = List.nth db (n-1) in + if Int.equal n' 0 then Tunknown else Tvar n') | Const (kn,u) -> let r = GlobRef.ConstRef kn in let typ = type_of env sg (EConstr.mkConstU (kn,u)) in (match flag_of_type env sg typ with - | (Logic,_) -> assert false (* Cf. logical cases above *) - | (Info, TypeScheme) -> + | (Logic,_) -> assert false (* Cf. logical cases above *) + | (Info, TypeScheme) -> let mlt = extract_type_app env sg db (r, type_sign env sg typ) args in (match (lookup_constant kn env).const_body with | Undef _ | OpaqueDef _ | Primitive _ -> mlt @@ -314,18 +314,18 @@ let rec extract_type env sg db j c args = | Def lbody -> let newc = applistc (get_body lbody) args in let mlt' = extract_type env sg db j newc [] in - (* ML type abbreviations interact badly with Coq *) - (* reduction, so [mlt] and [mlt'] might be different: *) - (* The more precise is [mlt'], extracted after reduction *) - (* The shortest is [mlt], which use abbreviations *) - (* If possible, we take [mlt], otherwise [mlt']. *) - if eq_ml_type (expand env mlt) (expand env mlt') then mlt else mlt') - | (Info, Default) -> + (* ML type abbreviations interact badly with Coq *) + (* reduction, so [mlt] and [mlt'] might be different: *) + (* The more precise is [mlt'], extracted after reduction *) + (* The shortest is [mlt], which use abbreviations *) + (* If possible, we take [mlt], otherwise [mlt']. *) + if eq_ml_type (expand env mlt) (expand env mlt') then mlt else mlt') + | (Info, Default) -> (* Not an ML type, for example [(c:forall X, X->X) Type nat] *) (match (lookup_constant kn env).const_body with | Undef _ | OpaqueDef _ | Primitive _ -> Tunknown (* Brutal approx ... *) - | Def lbody -> - (* We try to reduce. *) + | Def lbody -> + (* We try to reduce. *) let newc = applistc (get_body lbody) args in extract_type env sg db j newc [])) | Ind ((kn,i),u) -> @@ -415,15 +415,15 @@ and extract_really_ind env kn mib = (cf Vector and bug #2570) *) let equiv = if lang () != Ocaml || - (not (modular ()) && at_toplevel (MutInd.modpath kn)) || - KerName.equal (MutInd.canonical kn) (MutInd.user kn) + (not (modular ()) && at_toplevel (MutInd.modpath kn)) || + KerName.equal (MutInd.canonical kn) (MutInd.user kn) then - NoEquiv + NoEquiv else - begin - ignore (extract_ind env (MutInd.make1 (MutInd.canonical kn))); - Equiv (MutInd.canonical kn) - end + begin + ignore (extract_ind env (MutInd.make1 (MutInd.canonical kn))); + Equiv (MutInd.canonical kn) + end in (* Everything concerning parameters. *) (* We do that first, since they are common to all the [mib]. *) @@ -435,20 +435,20 @@ and extract_really_ind env kn mib = (* their type var list. *) let packets = Array.mapi - (fun i mip -> + (fun i mip -> let (_,u),_ = UnivGen.fresh_inductive_instance env (kn,i) in - let ar = Inductive.type_of_inductive env ((mib,mip),u) in + let ar = Inductive.type_of_inductive env ((mib,mip),u) in let ar = EConstr.of_constr ar in let info = (fst (flag_of_type env sg ar) = Info) in let s,v = if info then type_sign_vl env sg ar else [],[] in - let t = Array.make (Array.length mip.mind_nf_lc) [] in - { ip_typename = mip.mind_typename; - ip_consnames = mip.mind_consnames; - ip_logical = not info; - ip_sign = s; - ip_vars = v; - ip_types = t }, u) - mib.mind_packets + let t = Array.make (Array.length mip.mind_nf_lc) [] in + { ip_typename = mip.mind_typename; + ip_consnames = mip.mind_consnames; + ip_logical = not info; + ip_sign = s; + ip_vars = v; + ip_types = t }, u) + mib.mind_packets in add_ind kn mib @@ -461,85 +461,85 @@ and extract_really_ind env kn mib = for i = 0 to mib.mind_ntypes - 1 do let p,u = packets.(i) in if not p.ip_logical then - let types = arities_of_constructors env ((kn,i),u) in - for j = 0 to Array.length types - 1 do - let t = snd (decompose_prod_n npar types.(j)) in - let prods,head = dest_prod epar t in - let nprods = List.length prods in + let types = arities_of_constructors env ((kn,i),u) in + for j = 0 to Array.length types - 1 do + let t = snd (decompose_prod_n npar types.(j)) in + let prods,head = dest_prod epar t in + let nprods = List.length prods in let args = match Constr.kind head with | App (f,args) -> args (* [Constr.kind f = Ind ip] *) | _ -> [||] in - let dbmap = parse_ind_args p.ip_sign args (nprods + npar) in - let db = db_from_ind dbmap npar in + let dbmap = parse_ind_args p.ip_sign args (nprods + npar) in + let db = db_from_ind dbmap npar in p.ip_types.(j) <- extract_type_cons epar sg db dbmap (EConstr.of_constr t) (npar+1) - done + done done; (* Third pass: we determine special cases. *) let ind_info = try - let ip = (kn, 0) in + let ip = (kn, 0) in let r = GlobRef.IndRef ip in - if is_custom r then raise (I Standard); + if is_custom r then raise (I Standard); if mib.mind_finite == CoFinite then raise (I Coinductive); - if not (Int.equal mib.mind_ntypes 1) then raise (I Standard); - let p,u = packets.(0) in - if p.ip_logical then raise (I Standard); - if not (Int.equal (Array.length p.ip_types) 1) then raise (I Standard); - let typ = p.ip_types.(0) in - let l = List.filter (fun t -> not (isTdummy (expand env t))) typ in - if not (keep_singleton ()) && - Int.equal (List.length l) 1 && not (type_mem_kn kn (List.hd l)) - then raise (I Singleton); - if List.is_empty l then raise (I Standard); + if not (Int.equal mib.mind_ntypes 1) then raise (I Standard); + let p,u = packets.(0) in + if p.ip_logical then raise (I Standard); + if not (Int.equal (Array.length p.ip_types) 1) then raise (I Standard); + let typ = p.ip_types.(0) in + let l = List.filter (fun t -> not (isTdummy (expand env t))) typ in + if not (keep_singleton ()) && + Int.equal (List.length l) 1 && not (type_mem_kn kn (List.hd l)) + then raise (I Singleton); + if List.is_empty l then raise (I Standard); if mib.mind_record == Declarations.NotRecord then raise (I Standard); - (* Now we're sure it's a record. *) - (* First, we find its field names. *) - let rec names_prod t = match Constr.kind t with + (* Now we're sure it's a record. *) + (* First, we find its field names. *) + let rec names_prod t = match Constr.kind t with | Prod(n,_,t) -> n::(names_prod t) | LetIn(_,_,_,t) -> names_prod t - | Cast(t,_,_) -> names_prod t - | _ -> [] - in - let field_names = - List.skipn mib.mind_nparams (names_prod mip0.mind_user_lc.(0)) in - assert (Int.equal (List.length field_names) (List.length typ)); - let projs = ref Cset.empty in - let mp = MutInd.modpath kn in - let rec select_fields l typs = match l,typs with - | [],[] -> [] - | _::l, typ::typs when isTdummy (expand env typ) -> - select_fields l typs + | Cast(t,_,_) -> names_prod t + | _ -> [] + in + let field_names = + List.skipn mib.mind_nparams (names_prod mip0.mind_user_lc.(0)) in + assert (Int.equal (List.length field_names) (List.length typ)); + let projs = ref Cset.empty in + let mp = MutInd.modpath kn in + let rec select_fields l typs = match l,typs with + | [],[] -> [] + | _::l, typ::typs when isTdummy (expand env typ) -> + select_fields l typs | {binder_name=Anonymous}::l, typ::typs -> - None :: (select_fields l typs) + None :: (select_fields l typs) | {binder_name=Name id}::l, typ::typs -> - let knp = Constant.make2 mp (Label.of_id id) in - (* Is it safe to use [id] for projections [foo.id] ? *) - if List.for_all ((==) Keep) (type2signature env typ) - then projs := Cset.add knp !projs; + let knp = Constant.make2 mp (Label.of_id id) in + (* Is it safe to use [id] for projections [foo.id] ? *) + if List.for_all ((==) Keep) (type2signature env typ) + then projs := Cset.add knp !projs; Some (GlobRef.ConstRef knp) :: (select_fields l typs) - | _ -> assert false - in - let field_glob = select_fields field_names typ - in - (* Is this record officially declared with its projections ? *) - (* If so, we use this information. *) - begin try + | _ -> assert false + in + let field_glob = select_fields field_names typ + in + (* Is this record officially declared with its projections ? *) + (* If so, we use this information. *) + begin try let ty = Inductive.type_of_inductive env ((mib,mip0),u) in let n = nb_default_params env sg (EConstr.of_constr ty) in let check_proj kn = if Cset.mem kn !projs then add_projection n kn ip in - List.iter (Option.iter check_proj) (lookup_projections ip) - with Not_found -> () - end; - Record field_glob + List.iter (Option.iter check_proj) (lookup_projections ip) + with Not_found -> () + end; + Record field_glob with (I info) -> info in let i = {ind_kind = ind_info; - ind_nparams = npar; - ind_packets = Array.map fst packets; - ind_equiv = equiv } + ind_nparams = npar; + ind_packets = Array.map fst packets; + ind_equiv = equiv } in add_ind kn mib i; add_inductive_kind kn i.ind_kind; @@ -622,42 +622,42 @@ let rec extract_term env sg mle mlt c args = | Lambda (n, t, d) -> let id = map_annot id_of_name n in let idna = map_annot Name.mk_name id in - (match args with - | a :: l -> - (* We make as many [LetIn] as possible. *) + (match args with + | a :: l -> + (* We make as many [LetIn] as possible. *) let l' = List.map (EConstr.Vars.lift 1) l in let d' = EConstr.mkLetIn (idna,a,t,applistc d l') in extract_term env sg mle mlt d' [] | [] -> let env' = push_rel_assum (idna, t) env in - let id, a = + let id, a = try check_default env sg t; Id id.binder_name, new_meta() with NotDefault d -> Dummy, Tdummy d - in - let b = new_meta () in - (* If [mlt] cannot be unified with an arrow type, then magic! *) - let magic = needs_magic (mlt, Tarr (a, b)) in + in + let b = new_meta () in + (* If [mlt] cannot be unified with an arrow type, then magic! *) + let magic = needs_magic (mlt, Tarr (a, b)) in let d' = extract_term env' sg (Mlenv.push_type mle a) b d [] in - put_magic_if magic (MLlam (id, d'))) + put_magic_if magic (MLlam (id, d'))) | LetIn (n, c1, t1, c2) -> let id = map_annot id_of_name n in let env' = EConstr.push_rel (LocalDef (map_annot Name.mk_name id, c1, t1)) env in (* We directly push the args inside the [LetIn]. TODO: the opt_let_app flag is supposed to prevent that *) let args' = List.map (EConstr.Vars.lift 1) args in - (try + (try check_default env sg t1; - let a = new_meta () in + let a = new_meta () in let c1' = extract_term env sg mle a c1 [] in - (* The type of [c1'] is generalized and stored in [mle]. *) - let mle' = - if generalizable c1' - then Mlenv.push_gen mle a - else Mlenv.push_type mle a - in + (* The type of [c1'] is generalized and stored in [mle]. *) + let mle' = + if generalizable c1' + then Mlenv.push_gen mle a + else Mlenv.push_type mle a + in MLletin (Id id.binder_name, c1', extract_term env' sg mle' mlt c2 args') - with NotDefault d -> - let mle' = Mlenv.push_std_type mle (Tdummy d) in + with NotDefault d -> + let mle' = Mlenv.push_std_type mle (Tdummy d) in ast_pop (extract_term env' sg mle' mlt c2 args')) | Const (kn,_) -> extract_cst_app env sg mle mlt kn args @@ -667,9 +667,9 @@ let rec extract_term env sg mle mlt c args = let term = Retyping.expand_projection env (Evd.from_env env) p c [] in extract_term env sg mle mlt term args | Rel n -> - (* As soon as the expected [mlt] for the head is known, *) - (* we unify it with an fresh copy of the stored type of [Rel n]. *) - let extract_rel mlt = put_magic (mlt, Mlenv.get mle n) (MLrel n) + (* As soon as the expected [mlt] for the head is known, *) + (* we unify it with an fresh copy of the stored type of [Rel n]. *) + let extract_rel mlt = put_magic (mlt, Mlenv.get mle n) (MLrel n) in extract_app env sg mle mlt extract_rel args | Case ({ci_ind=ip},_,c0,br) -> extract_app env sg mle mlt (extract_case env sg mle (ip,c0,br)) args @@ -816,8 +816,8 @@ and extract_cons_app env sg mle mlt (((kn,i) as ip,j) as cp) args = put_magic_if magic1 (List.hd mla) (* assert (List.length mla = 1) *) else let typeargs = match snd (type_decomp type_cons) with - | Tglob (_,l) -> List.map type_simpl l - | _ -> assert false + | Tglob (_,l) -> List.map type_simpl l + | _ -> assert false in let typ = Tglob(GlobRef.IndRef ip, typeargs) in put_magic_if magic1 (MLcons (typ, GlobRef.ConstructRef cp, mla)) @@ -854,14 +854,14 @@ and extract_case env sg mle ((kn,i) as ip,c,br) mlt = (* The only non-informative case: [c] is of sort [Prop] *) if (sort_of env sg t) == InProp then begin - add_recursors env kn; (* May have passed unseen if logical ... *) - (* Logical singleton case: *) - (* [match c with C i j k -> t] becomes [t'] *) - assert (Int.equal br_size 1); - let s = iterate (fun l -> Kill Kprop :: l) ni.(0) [] in - let mlt = iterate (fun t -> Tarr (Tdummy Kprop, t)) ni.(0) mlt in + add_recursors env kn; (* May have passed unseen if logical ... *) + (* Logical singleton case: *) + (* [match c with C i j k -> t] becomes [t'] *) + assert (Int.equal br_size 1); + let s = iterate (fun l -> Kill Kprop :: l) ni.(0) [] in + let mlt = iterate (fun t -> Tarr (Tdummy Kprop, t)) ni.(0) mlt in let e = extract_maybe_term env sg mle mlt br.(0) in - snd (case_expunge s e) + snd (case_expunge s e) end else let mi = extract_ind env kn in @@ -873,32 +873,32 @@ and extract_case env sg mle ((kn,i) as ip,c,br) mlt = (* The extraction of each branch. *) let extract_branch i = let r = GlobRef.ConstructRef (ip,i+1) in - (* The types of the arguments of the corresponding constructor. *) - let f t = type_subst_vect metas (expand env t) in - let l = List.map f oi.ip_types.(i) in - (* the corresponding signature *) - let s = List.map (type2sign env) oi.ip_types.(i) in - let s = sign_with_implicits r s mi.ind_nparams in - (* Extraction of the branch (in functional form). *) + (* The types of the arguments of the corresponding constructor. *) + let f t = type_subst_vect metas (expand env t) in + let l = List.map f oi.ip_types.(i) in + (* the corresponding signature *) + let s = List.map (type2sign env) oi.ip_types.(i) in + let s = sign_with_implicits r s mi.ind_nparams in + (* Extraction of the branch (in functional form). *) let e = extract_maybe_term env sg mle (type_recomp (l,mlt)) br.(i) in - (* We suppress dummy arguments according to signature. *) - let ids,e = case_expunge s e in - (List.rev ids, Pusual r, e) + (* We suppress dummy arguments according to signature. *) + let ids,e = case_expunge s e in + (List.rev ids, Pusual r, e) in if mi.ind_kind == Singleton then - begin - (* Informative singleton case: *) - (* [match c with C i -> t] becomes [let i = c' in t'] *) - assert (Int.equal br_size 1); - let (ids,_,e') = extract_branch 0 in - assert (Int.equal (List.length ids) 1); - MLletin (tmp_id (List.hd ids),a,e') - end + begin + (* Informative singleton case: *) + (* [match c with C i -> t] becomes [let i = c' in t'] *) + assert (Int.equal br_size 1); + let (ids,_,e') = extract_branch 0 in + assert (Int.equal (List.length ids) 1); + MLletin (tmp_id (List.hd ids),a,e') + end else - (* Standard case: we apply [extract_branch]. *) - let typs = List.map type_simpl (Array.to_list metas) in + (* Standard case: we apply [extract_branch]. *) + let typs = List.map type_simpl (Array.to_list metas) in let typ = Tglob (GlobRef.IndRef ip,typs) in - MLcase (typ, a, Array.init br_size extract_branch) + MLcase (typ, a, Array.init br_size extract_branch) (*s Extraction of a (co)-fixpoint. *) @@ -932,7 +932,7 @@ let rec gentypvar_ok sg c = match EConstr.kind sg c with | Lambda _ | Const _ -> true | App (c,v) -> (* if all arguments are variables, these variables will - disappear after extraction (see [empty_s] below) *) + disappear after extraction (see [empty_s] below) *) Array.for_all (EConstr.isRel sg) v && gentypvar_ok sg c | Cast (c,_,_) -> gentypvar_ok sg c | _ -> false @@ -962,7 +962,7 @@ let extract_std_constant env sg kn body typ = else let s,s' = List.chop m s in if List.for_all ((==) Keep) s' && - (lang () == Haskell || sign_kind s != UnsafeLogicalSig) + (lang () == Haskell || sign_kind s != UnsafeLogicalSig) then decompose_lam_n sg m body else decomp_lams_eta_n n m env sg body typ in @@ -1114,27 +1114,27 @@ let extract_constant env kn cb = | (Info,TypeScheme) -> (match cb.const_body with | Primitive _ | Undef _ -> warn_info (); mk_typ_ax () - | Def c -> + | Def c -> (match Recordops.find_primitive_projection kn with | None -> mk_typ (get_body c) | Some p -> let body = fake_match_projection env p in mk_typ (EConstr.of_constr body)) - | OpaqueDef c -> - add_opaque r; + | OpaqueDef c -> + add_opaque r; if access_opaque () then mk_typ (get_opaque env c) else mk_typ_ax ()) | (Info,Default) -> (match cb.const_body with | Primitive _ | Undef _ -> warn_info (); mk_ax () - | Def c -> + | Def c -> (match Recordops.find_primitive_projection kn with | None -> mk_def (get_body c) | Some p -> let body = fake_match_projection env p in mk_def (EConstr.of_constr body)) - | OpaqueDef c -> - add_opaque r; + | OpaqueDef c -> + add_opaque r; if access_opaque () then mk_def (get_opaque env c) else mk_ax ()) with SingletonInductiveBecomesProp id -> @@ -1150,10 +1150,10 @@ let extract_constant_spec env kn cb = | (Logic, Default) -> Sval (r, Tdummy Kprop) | (Info, TypeScheme) -> let s,vl = type_sign_vl env sg typ in - (match cb.const_body with + (match cb.const_body with | Undef _ | OpaqueDef _ | Primitive _ -> Stype (r, vl, None) - | Def body -> - let db = db_from_sign s in + | Def body -> + let db = db_from_sign s in let body = get_body body in let t = extract_type_scheme env sg db body (List.length s) in Stype (r, vl, Some t)) @@ -1197,9 +1197,9 @@ let extract_inductive env kn = let rec filter i = function | [] -> [] | t::l -> - let l' = filter (succ i) l in - if isTdummy (expand env t) || Int.Set.mem i implicits then l' - else t::l' + let l' = filter (succ i) l in + if isTdummy (expand env t) || Int.Set.mem i implicits then l' + else t::l' in filter (1+ind.ind_nparams) l in let packets = diff --git a/plugins/extraction/haskell.ml b/plugins/extraction/haskell.ml index 4769bef475..f0053ba6b5 100644 --- a/plugins/extraction/haskell.ml +++ b/plugins/extraction/haskell.ml @@ -110,15 +110,15 @@ let rec pp_type par vl t = with Failure _ -> (str "a" ++ int i)) | Tglob (r,[]) -> pp_global Type r | Tglob (GlobRef.IndRef(kn,0),l) - when not (keep_singleton ()) && MutInd.equal kn (mk_ind "Coq.Init.Specif" "sig") -> - pp_type true vl (List.hd l) + when not (keep_singleton ()) && MutInd.equal kn (mk_ind "Coq.Init.Specif" "sig") -> + pp_type true vl (List.hd l) | Tglob (r,l) -> - pp_par par - (pp_global Type r ++ spc () ++ - prlist_with_sep spc (pp_type true vl) l) + pp_par par + (pp_global Type r ++ spc () ++ + prlist_with_sep spc (pp_type true vl) l) | Tarr (t1,t2) -> - pp_par par - (pp_rec true t1 ++ spc () ++ str "->" ++ spc () ++ pp_rec false t2) + pp_par par + (pp_rec true t1 ++ spc () ++ str "->" ++ spc () ++ pp_rec false t2) | Tdummy _ -> str "()" | Tunknown -> str "Any" | Taxiom -> str "() -- AXIOM TO BE REALIZED" ++ fnl () @@ -141,77 +141,77 @@ let rec pp_expr par env args = and apply2 st = pp_apply2 st par args in function | MLrel n -> - let id = get_db_name n env in + let id = get_db_name n env in (* Try to survive to the occurrence of a Dummy rel. TODO: we should get rid of this hack (cf. BZ#592) *) let id = if Id.equal id dummy_name then Id.of_string "__" else id in apply (Id.print id) | MLapp (f,args') -> - let stl = List.map (pp_expr true env []) args' in + let stl = List.map (pp_expr true env []) args' in pp_expr par env (stl @ args) f | MLlam _ as a -> - let fl,a' = collect_lams a in - let fl,env' = push_vars (List.map id_of_mlid fl) env in - let st = (pp_abst (List.rev fl) ++ pp_expr false env' [] a') in - apply2 st + let fl,a' = collect_lams a in + let fl,env' = push_vars (List.map id_of_mlid fl) env in + let st = (pp_abst (List.rev fl) ++ pp_expr false env' [] a') in + apply2 st | MLletin (id,a1,a2) -> - let i,env' = push_vars [id_of_mlid id] env in - let pp_id = Id.print (List.hd i) - and pp_a1 = pp_expr false env [] a1 - and pp_a2 = pp_expr (not par && expr_needs_par a2) env' [] a2 in - let pp_def = - str "let {" ++ cut () ++ - hov 1 (pp_id ++ str " = " ++ pp_a1 ++ str "}") - in - apply2 (hv 0 (hv 0 (hv 1 pp_def ++ spc () ++ str "in") ++ - spc () ++ hov 0 pp_a2)) + let i,env' = push_vars [id_of_mlid id] env in + let pp_id = Id.print (List.hd i) + and pp_a1 = pp_expr false env [] a1 + and pp_a2 = pp_expr (not par && expr_needs_par a2) env' [] a2 in + let pp_def = + str "let {" ++ cut () ++ + hov 1 (pp_id ++ str " = " ++ pp_a1 ++ str "}") + in + apply2 (hv 0 (hv 0 (hv 1 pp_def ++ spc () ++ str "in") ++ + spc () ++ hov 0 pp_a2)) | MLglob r -> - apply (pp_global Term r) + apply (pp_global Term r) | MLcons (_,r,a) as c -> assert (List.is_empty args); begin match a with - | _ when is_native_char c -> pp_native_char c - | [] -> pp_global Cons r - | [a] -> - pp_par par (pp_global Cons r ++ spc () ++ pp_expr true env [] a) - | _ -> - pp_par par (pp_global Cons r ++ spc () ++ - prlist_with_sep spc (pp_expr true env []) a) - end + | _ when is_native_char c -> pp_native_char c + | [] -> pp_global Cons r + | [a] -> + pp_par par (pp_global Cons r ++ spc () ++ pp_expr true env [] a) + | _ -> + pp_par par (pp_global Cons r ++ spc () ++ + prlist_with_sep spc (pp_expr true env []) a) + end | MLtuple l -> assert (List.is_empty args); pp_boxed_tuple (pp_expr true env []) l | MLcase (_,t, pv) when is_custom_match pv -> if not (is_regular_match pv) then - user_err Pp.(str "Cannot mix yet user-given match and general patterns."); - let mkfun (ids,_,e) = - if not (List.is_empty ids) then named_lams (List.rev ids) e - else dummy_lams (ast_lift 1 e) 1 - in - let pp_branch tr = pp_expr true env [] (mkfun tr) ++ fnl () in - let inner = - str (find_custom_match pv) ++ fnl () ++ - prvect pp_branch pv ++ - pp_expr true env [] t - in - apply2 (hov 2 inner) + user_err Pp.(str "Cannot mix yet user-given match and general patterns."); + let mkfun (ids,_,e) = + if not (List.is_empty ids) then named_lams (List.rev ids) e + else dummy_lams (ast_lift 1 e) 1 + in + let pp_branch tr = pp_expr true env [] (mkfun tr) ++ fnl () in + let inner = + str (find_custom_match pv) ++ fnl () ++ + prvect pp_branch pv ++ + pp_expr true env [] t + in + apply2 (hov 2 inner) | MLcase (typ,t,pv) -> apply2 - (v 0 (str "case " ++ pp_expr false env [] t ++ str " of {" ++ - fnl () ++ pp_pat env pv)) + (v 0 (str "case " ++ pp_expr false env [] t ++ str " of {" ++ + fnl () ++ pp_pat env pv)) | MLfix (i,ids,defs) -> - let ids',env' = push_vars (List.rev (Array.to_list ids)) env in - pp_fix par env' i (Array.of_list (List.rev ids'),defs) args + let ids',env' = push_vars (List.rev (Array.to_list ids)) env in + pp_fix par env' i (Array.of_list (List.rev ids'),defs) args | MLexn s -> - (* An [MLexn] may be applied, but I don't really care. *) - pp_par par (str "Prelude.error" ++ spc () ++ qs s) + (* An [MLexn] may be applied, but I don't really care. *) + pp_par par (str "Prelude.error" ++ spc () ++ qs s) | MLdummy k -> (* An [MLdummy] may be applied, but I don't really care. *) (match msg_of_implicit k with | "" -> str "__" | s -> str "__" ++ spc () ++ pp_bracket_comment (str s)) | MLmagic a -> - pp_apply (str "unsafeCoerce") par (pp_expr true env [] a :: args) + pp_apply (str "unsafeCoerce") par (pp_expr true env [] a :: args) | MLaxiom -> pp_par par (str "Prelude.error \"AXIOM TO BE REALIZED\"") | MLuint _ -> pp_par par (str "Prelude.error \"EXTRACTION OF UINT NOT IMPLEMENTED\"") @@ -232,16 +232,16 @@ and pp_gen_pat par ids env = function and pp_one_pat env (ids,p,t) = let ids',env' = push_vars (List.rev_map id_of_mlid ids) env in hov 2 (str " " ++ - pp_gen_pat false (List.rev ids') env' p ++ - str " ->" ++ spc () ++ - pp_expr (expr_needs_par t) env' [] t) + pp_gen_pat false (List.rev ids') env' p ++ + str " ->" ++ spc () ++ + pp_expr (expr_needs_par t) env' [] t) and pp_pat env pv = prvecti (fun i x -> pp_one_pat env pv.(i) ++ if Int.equal i (Array.length pv - 1) then str "}" else - (str ";" ++ fnl ())) + (str ";" ++ fnl ())) pv (*s names of the functions ([ids]) are already pushed in [env], @@ -251,10 +251,10 @@ and pp_fix par env i (ids,bl) args = pp_par par (v 0 (v 1 (str "let {" ++ fnl () ++ - prvect_with_sep (fun () -> str ";" ++ fnl ()) - (fun (fi,ti) -> pp_function env (Id.print fi) ti) - (Array.map2 (fun a b -> a,b) ids bl) ++ - str "}") ++ + prvect_with_sep (fun () -> str ";" ++ fnl ()) + (fun (fi,ti) -> pp_function env (Id.print fi) ti) + (Array.map2 (fun a b -> a,b) ids bl) ++ + str "}") ++ fnl () ++ str "in " ++ pp_apply (Id.print ids.(i)) false args)) and pp_function env f t = @@ -269,17 +269,17 @@ and pp_function env f t = let pp_logical_ind packet = pp_comment (Id.print packet.ip_typename ++ str " : logical inductive") ++ pp_comment (str "with constructors : " ++ - prvect_with_sep spc Id.print packet.ip_consnames) + prvect_with_sep spc Id.print packet.ip_consnames) let pp_singleton kn packet = let name = pp_global Type (GlobRef.IndRef (kn,0)) in let l = rename_tvars keywords packet.ip_vars in hov 2 (str "type " ++ name ++ spc () ++ - prlist_with_sep spc Id.print l ++ - (if not (List.is_empty l) then str " " else mt ()) ++ str "=" ++ spc () ++ - pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++ - pp_comment (str "singleton inductive, whose constructor was " ++ - Id.print packet.ip_consnames.(0))) + prlist_with_sep spc Id.print l ++ + (if not (List.is_empty l) then str " " else mt ()) ++ str "=" ++ spc () ++ + pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++ + pp_comment (str "singleton inductive, whose constructor was " ++ + Id.print packet.ip_consnames.(0))) let pp_one_ind ip pl cv = let pl = rename_tvars keywords pl in @@ -288,8 +288,8 @@ let pp_one_ind ip pl cv = match l with | [] -> (mt ()) | _ -> (str " " ++ - prlist_with_sep - (fun () -> (str " ")) (pp_type true pl) l)) + prlist_with_sep + (fun () -> (str " ")) (pp_type true pl) l)) in str (if Array.is_empty cv then "type " else "data ") ++ pp_global Type (GlobRef.IndRef ip) ++ @@ -298,7 +298,7 @@ let pp_one_ind ip pl cv = else (fnl () ++ str " " ++ v 0 (str " " ++ - prvect_with_sep (fun () -> fnl () ++ str "| ") pp_constructor + prvect_with_sep (fun () -> fnl () ++ str "| ") pp_constructor (Array.mapi (fun i c -> GlobRef.ConstructRef (ip,i+1),c) cv))) let rec pp_ind first kn i ind = @@ -310,10 +310,10 @@ let rec pp_ind first kn i ind = if is_custom (GlobRef.IndRef (kn,i)) then pp_ind first kn (i+1) ind else if p.ip_logical then - pp_logical_ind p ++ pp_ind first kn (i+1) ind + pp_logical_ind p ++ pp_ind first kn (i+1) ind else - pp_one_ind ip p.ip_vars p.ip_types ++ fnl () ++ - pp_ind false kn (i+1) ind + pp_one_ind ip p.ip_vars p.ip_types ++ fnl () ++ + pp_ind false kn (i+1) ind (*s Pretty-printing of a declaration. *) @@ -325,45 +325,45 @@ let pp_decl = function | Dtype (r, l, t) -> if is_inline_custom r then mt () else - let l = rename_tvars keywords l in - let st = - try - let ids,s = find_type_custom r in - prlist (fun id -> str (id^" ")) ids ++ str "=" ++ spc () ++ str s - with Not_found -> - prlist (fun id -> Id.print id ++ str " ") l ++ - if t == Taxiom then str "= () -- AXIOM TO BE REALIZED" ++ fnl () - else str "=" ++ spc () ++ pp_type false l t - in - hov 2 (str "type " ++ pp_global Type r ++ spc () ++ st) ++ fnl2 () + let l = rename_tvars keywords l in + let st = + try + let ids,s = find_type_custom r in + prlist (fun id -> str (id^" ")) ids ++ str "=" ++ spc () ++ str s + with Not_found -> + prlist (fun id -> Id.print id ++ str " ") l ++ + if t == Taxiom then str "= () -- AXIOM TO BE REALIZED" ++ fnl () + else str "=" ++ spc () ++ pp_type false l t + in + hov 2 (str "type " ++ pp_global Type r ++ spc () ++ st) ++ fnl2 () | Dfix (rv, defs, typs) -> let names = Array.map - (fun r -> if is_inline_custom r then mt () else pp_global Term r) rv + (fun r -> if is_inline_custom r then mt () else pp_global Term r) rv in prvecti - (fun i r -> - let void = is_inline_custom r || - (not (is_custom r) && + (fun i r -> + let void = is_inline_custom r || + (not (is_custom r) && match defs.(i) with MLexn "UNUSED" -> true | _ -> false) - in - if void then mt () - else - hov 2 (names.(i) ++ str " :: " ++ pp_type false [] typs.(i)) ++ fnl () ++ - (if is_custom r then - (names.(i) ++ str " = " ++ str (find_custom r)) - else - (pp_function (empty_env ()) names.(i) defs.(i))) - ++ fnl2 ()) - rv + in + if void then mt () + else + hov 2 (names.(i) ++ str " :: " ++ pp_type false [] typs.(i)) ++ fnl () ++ + (if is_custom r then + (names.(i) ++ str " = " ++ str (find_custom r)) + else + (pp_function (empty_env ()) names.(i) defs.(i))) + ++ fnl2 ()) + rv | Dterm (r, a, t) -> if is_inline_custom r then mt () else - let e = pp_global Term r in - hov 2 (e ++ str " :: " ++ pp_type false [] t) ++ fnl () ++ - if is_custom r then - hov 0 (e ++ str " = " ++ str (find_custom r) ++ fnl2 ()) - else - hov 0 (pp_function (empty_env ()) e a ++ fnl2 ()) + let e = pp_global Term r in + hov 2 (e ++ str " :: " ++ pp_type false [] t) ++ fnl () ++ + if is_custom r then + hov 0 (e ++ str " = " ++ str (find_custom r) ++ fnl2 ()) + else + hov 0 (pp_function (empty_env ()) e a ++ fnl2 ()) let rec pp_structure_elem = function | (l,SEdecl d) -> pp_decl d diff --git a/plugins/extraction/mlutil.ml b/plugins/extraction/mlutil.ml index 44b95ae4c1..fc0ba95b98 100644 --- a/plugins/extraction/mlutil.ml +++ b/plugins/extraction/mlutil.ml @@ -200,10 +200,10 @@ module Mlenv = struct let rec meta2var t = match t with | Tmeta {contents=Some u} -> meta2var u | Tmeta ({id=i} as m) -> - (try Tvar (Int.Map.find i !map) - with Not_found -> - if Metaset.mem m mle.free then t - else Tvar (add_new i)) + (try Tvar (Int.Map.find i !map) + with Not_found -> + if Metaset.mem m mle.free then t + else Tvar (add_new i)) | Tarr (t1,t2) -> Tarr (meta2var t1, meta2var t2) | Tglob (r,l) -> Tglob (r, List.map meta2var l) | t -> t @@ -279,9 +279,9 @@ let type_expand env t = let rec expand = function | Tmeta {contents = Some t} -> expand t | Tglob (r,l) -> - (match env r with - | Some mlt -> expand (type_subst_list l mlt) - | None -> Tglob (r, List.map expand l)) + (match env r with + | Some mlt -> expand (type_subst_list l mlt) + | None -> Tglob (r, List.map expand l)) | Tarr (a,b) -> Tarr (expand a, expand b) | a -> a in if Table.type_expand () then expand t else t @@ -348,8 +348,8 @@ let type_expunge_from_sign env s t = | _, Tmeta {contents = Some t} -> expunge s t | _, Tglob (r,l) -> (match env r with - | Some mlt -> expunge s (type_subst_list l mlt) - | None -> assert false) + | Some mlt -> expunge s (type_subst_list l mlt) + | None -> assert false) | _ -> assert false in let t = expunge (sign_no_final_keeps s) t in @@ -426,7 +426,7 @@ let ast_iter_rel f = | MLlam (_,a) -> iter (n+1) a | MLletin (_,a,b) -> iter n a; iter (n+1) b | MLcase (_,a,v) -> - iter n a; Array.iter (fun (l,_,t) -> iter (n + (List.length l)) t) v + iter n a; Array.iter (fun (l,_,t) -> iter (n + (List.length l)) t) v | MLfix (_,ids,v) -> let k = Array.length ids in Array.iter (iter (n+k)) v | MLapp (a,l) -> iter n a; List.iter (iter n) l | MLcons (_,_,l) | MLtuple l -> List.iter (iter n) l @@ -512,8 +512,8 @@ let nb_occur_match = | MLrel i -> if Int.equal i k then 1 else 0 | MLcase(_,a,v) -> (nb k a) + - Array.fold_left - (fun r (ids,_,a) -> max r (nb (k+(List.length ids)) a)) 0 v + Array.fold_left + (fun r (ids,_,a) -> max r (nb (k+(List.length ids)) a)) 0 v | MLletin (_,a,b) -> (nb k a) + (nb (k+1) b) | MLfix (_,ids,v) -> let k = k+(Array.length ids) in Array.fold_left (fun r a -> r+(nb k a)) 0 v @@ -605,10 +605,10 @@ let ast_pop t = ast_lift (-1) t let permut_rels k k' = let rec permut n = function | MLrel i as a -> - let i' = i-n in - if i'<1 || i'>k+k' then a - else if i'<=k then MLrel (i+k') - else MLrel (i-k) + let i' = i-n in + if i'<1 || i'>k+k' then a + else if i'<=k then MLrel (i+k') + else MLrel (i-k) | a -> ast_map_lift permut n a in permut 0 @@ -618,10 +618,10 @@ let permut_rels k k' = let ast_subst e = let rec subst n = function | MLrel i as a -> - let i' = i-n in - if Int.equal i' 1 then ast_lift n e - else if i'<1 then a - else MLrel (i-1) + let i' = i-n in + if Int.equal i' 1 then ast_lift n e + else if i'<1 then a + else MLrel (i-1) | a -> ast_map_lift subst n a in subst 0 @@ -633,13 +633,13 @@ let ast_subst e = let gen_subst v d t = let rec subst n = function | MLrel i as a -> - let i'= i-n in - if i' < 1 then a - else if i' <= Array.length v then - match v.(i'-1) with - | None -> assert false - | Some u -> ast_lift n u - else MLrel (i+d) + let i'= i-n in + if i' < 1 then a + else if i' <= Array.length v then + match v.(i'-1) with + | None -> assert false + | Some u -> ast_lift n u + else MLrel (i+d) | a -> ast_map_lift subst n a in subst 0 t @@ -661,18 +661,18 @@ let is_regular_match br = else try let get_r (ids,pat,c) = - match pat with - | Pusual r -> r - | Pcons (r,l) -> + match pat with + | Pusual r -> r + | Pcons (r,l) -> let is_rel i = function Prel j -> Int.equal i j | _ -> false in - if not (List.for_all_i is_rel 1 (List.rev l)) - then raise Impossible; - r - | _ -> raise Impossible + if not (List.for_all_i is_rel 1 (List.rev l)) + then raise Impossible; + r + | _ -> raise Impossible in let ind = match get_r br.(0) with | GlobRef.ConstructRef (ind,_) -> ind - | _ -> raise Impossible + | _ -> raise Impossible in let is_ref i tr = match get_r tr with | GlobRef.ConstructRef (ind', j) -> eq_ind ind ind' && Int.equal j (i + 1) @@ -767,20 +767,20 @@ let eta_red e = if Int.equal n 0 then e else match t with | MLapp (f,a) -> - let m = List.length a in - let ids,body,args = - if Int.equal m n then - [], f, a - else if m < n then - List.skipn m ids, f, a - else (* m > n *) - let a1,a2 = List.chop (m-n) a in - [], MLapp (f,a1), a2 - in - let p = List.length args in - if test_eta_args_lift 0 p args && not (ast_occurs_itvl 1 p body) - then named_lams ids (ast_lift (-p) body) - else e + let m = List.length a in + let ids,body,args = + if Int.equal m n then + [], f, a + else if m < n then + List.skipn m ids, f, a + else (* m > n *) + let a1,a2 = List.chop (m-n) a in + [], MLapp (f,a1), a2 + in + let p = List.length args in + if test_eta_args_lift 0 p args && not (ast_occurs_itvl 1 p body) + then named_lams ids (ast_lift (-p) body) + else e | _ -> e (* Performs an eta-reduction when the core is atomic and value, @@ -804,11 +804,11 @@ let rec linear_beta_red a t = match a,t with | [], _ -> t | a0::a, MLlam (id,t) -> (match nb_occur_match t with - | 0 -> linear_beta_red a (ast_pop t) - | 1 -> linear_beta_red a (ast_subst a0 t) - | _ -> - let a = List.map (ast_lift 1) a in - MLletin (id, a0, linear_beta_red a t)) + | 0 -> linear_beta_red a (ast_pop t) + | 1 -> linear_beta_red a (ast_subst a0 t) + | _ -> + let a = List.map (ast_lift 1) a in + MLletin (id, a0, linear_beta_red a t)) | _ -> MLapp (t, a) let rec tmp_head_lams = function @@ -860,10 +860,10 @@ let branch_as_fun typ (l,p,c) = in let rec genrec n = function | MLrel i as c -> - let i' = i-n in - if i'<1 then c - else if i'>nargs then MLrel (i-nargs+1) - else raise Impossible + let i' = i-n in + if i'<1 then c + else if i'>nargs then MLrel (i-nargs+1) + else raise Impossible | MLcons _ as cons' when eq_ml_ast cons' (ast_lift n cons) -> MLrel (n+1) | a -> ast_map_lift genrec n a in genrec 0 c @@ -909,8 +909,8 @@ let census_add, census_max, census_clean = let len = ref 0 and lst = ref Int.Set.empty and elm = ref MLaxiom in List.iter (fun (e, s) -> - let n = Int.Set.cardinal s in - if n > !len then begin len := n; lst := s; elm := e end) + let n = Int.Set.cardinal s in + if n > !len then begin len := n; lst := s; elm := e end) !h; (!elm,!lst) in @@ -931,9 +931,9 @@ let factor_branches o typ br = census_clean (); for i = 0 to Array.length br - 1 do if o.opt_case_idr then - (try census_add (branch_as_fun typ br.(i)) i with Impossible -> ()); + (try census_add (branch_as_fun typ br.(i)) i with Impossible -> ()); if o.opt_case_cst then - (try census_add (branch_as_cst br.(i)) i with Impossible -> ()); + (try census_add (branch_as_cst br.(i)) i with Impossible -> ()); done; let br_factor, br_set = census_max () in census_clean (); @@ -956,9 +956,9 @@ let is_exn = function MLexn _ -> true | _ -> false let permut_case_fun br acc = let nb = ref max_int in Array.iter (fun (_,_,t) -> - let ids, c = collect_lams t in - let n = List.length ids in - if (n < !nb) && (not (is_exn c)) then nb := n) br; + let ids, c = collect_lams t in + let n = List.length ids in + if (n < !nb) && (not (is_exn c)) then nb := n) br; if Int.equal !nb max_int || Int.equal !nb 0 then ([],br) else begin let br = Array.copy br in @@ -967,11 +967,11 @@ let permut_case_fun br acc = let (l,p,t) = br.(i) in let local_nb = nb_lams t in if local_nb < !nb then (* t = MLexn ... *) - br.(i) <- (l,p,remove_n_lams local_nb t) + br.(i) <- (l,p,remove_n_lams local_nb t) else begin - let local_ids,t = collect_n_lams !nb t in - ids := merge_ids !ids local_ids; - br.(i) <- (l,p,permut_rels !nb (List.length l) t) + let local_ids,t = collect_n_lams !nb t in + ids := merge_ids !ids local_ids; + br.(i) <- (l,p,permut_rels !nb (List.length l) t) end done; (!ids,br) @@ -1011,9 +1011,9 @@ let iota_gen br hd = let rec iota k = function | MLcons (typ,r,a) -> iota_red 0 k br (typ,r,a) | MLcase(typ,e,br') -> - let new_br = - Array.map (fun (i,p,c)->(i,p,iota (k+(List.length i)) c)) br' - in MLcase(typ,e,new_br) + let new_br = + Array.map (fun (i,p,c)->(i,p,iota (k+(List.length i)) c)) br' + in MLcase(typ,e,new_br) | _ -> raise Impossible in iota 0 hd @@ -1061,17 +1061,17 @@ let rec simpl o = function | MLletin(id,c,e) -> let e = simpl o e in if - (is_atomic c) || (is_atomic e) || - (let n = nb_occur_match e in - (Int.equal n 0 || (Int.equal n 1 && expand_linear_let o id e))) + (is_atomic c) || (is_atomic e) || + (let n = nb_occur_match e in + (Int.equal n 0 || (Int.equal n 1 && expand_linear_let o id e))) then - simpl o (ast_subst c e) + simpl o (ast_subst c e) else - MLletin(id, simpl o c, e) + MLletin(id, simpl o c, e) | MLfix(i,ids,c) -> let n = Array.length ids in if ast_occurs_itvl 1 n c.(i) then - MLfix (i, ids, Array.map (simpl o) c) + MLfix (i, ids, Array.map (simpl o) c) else simpl o (ast_lift (-n) c.(i)) (* Dummy fixpoint *) | MLmagic(MLmagic _ as e) -> simpl o e | MLmagic(MLapp (f,l)) -> simpl o (MLapp (MLmagic f, l)) @@ -1094,12 +1094,12 @@ and simpl_app o a = function simpl o (MLapp (ast_pop t, List.tl a)) | MLlam (id,t) -> (* Beta redex *) (match nb_occur_match t with - | 0 -> simpl o (MLapp (ast_pop t, List.tl a)) - | 1 when (is_tmp id || o.opt_lin_beta) -> - simpl o (MLapp (ast_subst (List.hd a) t, List.tl a)) - | _ -> - let a' = List.map (ast_lift 1) (List.tl a) in - simpl o (MLletin (id, List.hd a, MLapp (t, a')))) + | 0 -> simpl o (MLapp (ast_pop t, List.tl a)) + | 1 when (is_tmp id || o.opt_lin_beta) -> + simpl o (MLapp (ast_subst (List.hd a) t, List.tl a)) + | _ -> + let a' = List.map (ast_lift 1) (List.tl a) in + simpl o (MLletin (id, List.hd a, MLapp (t, a')))) | MLmagic (MLlam (id,t)) -> (* When we've at least one argument, we permute the magic and the lambda, to simplify things a bit (see #2795). @@ -1111,14 +1111,14 @@ and simpl_app o a = function | MLcase (typ,e,br) when o.opt_case_app -> (* Application of a case: we push arguments inside *) let br' = - Array.map - (fun (l,p,t) -> - let k = List.length l in - let a' = List.map (ast_lift k) a in - (l, p, simpl o (MLapp (t,a')))) br + Array.map + (fun (l,p,t) -> + let k = List.length l in + let a' = List.map (ast_lift k) a in + (l, p, simpl o (MLapp (t,a')))) br in simpl o (MLcase (typ,e,br')) | (MLdummy _ | MLexn _) as e -> e - (* We just discard arguments in those cases. *) + (* We just discard arguments in those cases. *) | f -> MLapp (f,a) (* Invariant : all empty matches should now be [MLexn] *) @@ -1139,19 +1139,19 @@ and simpl_case o typ br e = if lang() == Scheme || is_custom_match br then MLcase (typ, e, br) else match factor_branches o typ br with - | Some (f,ints) when Int.equal (Int.Set.cardinal ints) (Array.length br) -> - (* If all branches have been factorized, we remove the match *) - simpl o (MLletin (Tmp anonymous_name, e, f)) - | Some (f,ints) -> - let last_br = - if ast_occurs 1 f then ([Tmp anonymous_name], Prel 1, f) - else ([], Pwild, ast_pop f) - in - let brl = Array.to_list br in - let brl_opt = List.filteri (fun i _ -> not (Int.Set.mem i ints)) brl in - let brl_opt = brl_opt @ [last_br] in - MLcase (typ, e, Array.of_list brl_opt) - | None -> MLcase (typ, e, br) + | Some (f,ints) when Int.equal (Int.Set.cardinal ints) (Array.length br) -> + (* If all branches have been factorized, we remove the match *) + simpl o (MLletin (Tmp anonymous_name, e, f)) + | Some (f,ints) -> + let last_br = + if ast_occurs 1 f then ([Tmp anonymous_name], Prel 1, f) + else ([], Pwild, ast_pop f) + in + let brl = Array.to_list br in + let brl_opt = List.filteri (fun i _ -> not (Int.Set.mem i ints)) brl in + let brl_opt = brl_opt @ [last_br] in + MLcase (typ, e, Array.of_list brl_opt) + | None -> MLcase (typ, e, br) (*S Local prop elimination. *) (* We try to eliminate as many [prop] as possible inside an [ml_ast]. *) @@ -1230,8 +1230,8 @@ let kill_dummy_lams sign c = let eta_expansion_sign s (ids,c) = let rec abs ids rels i = function | [] -> - let a = List.rev_map (function MLrel x -> MLrel (i-x) | a -> a) rels - in ids, MLapp (ast_lift (i-1) c, a) + let a = List.rev_map (function MLrel x -> MLrel (i-x) | a -> a) rels + in ids, MLapp (ast_lift (i-1) c, a) | Keep :: l -> abs (anonymous :: ids) (MLrel i :: rels) (i+1) l | Kill k :: l -> abs (Dummy :: ids) (MLdummy k :: rels) (i+1) l in abs ids [] 1 s @@ -1275,14 +1275,14 @@ let kill_dummy_args (ids,bl) r t = in let rec killrec n = function | MLapp(e, a) when found n e -> - let k = max 0 (m - (List.length a)) in - let a = List.map (killrec n) a in - let a = List.map (ast_lift k) a in - let a = select_via_bl sign (a @ (eta_args k)) in - named_lams (List.firstn k ids) (MLapp (ast_lift k e, a)) + let k = max 0 (m - (List.length a)) in + let a = List.map (killrec n) a in + let a = List.map (ast_lift k) a in + let a = select_via_bl sign (a @ (eta_args k)) in + named_lams (List.firstn k ids) (MLapp (ast_lift k e, a)) | e when found n e -> - let a = select_via_bl sign (eta_args m) in - named_lams ids (MLapp (ast_lift m e, a)) + let a = select_via_bl sign (eta_args m) in + named_lams ids (MLapp (ast_lift m e, a)) | e -> ast_map_lift killrec n e in killrec 0 t @@ -1294,32 +1294,32 @@ let sign_of_args a = let rec kill_dummy = function | MLfix(i,fi,c) -> (try - let k,c = kill_dummy_fix i c [] in - ast_subst (MLfix (i,fi,c)) (kill_dummy_args k 1 (MLrel 1)) + let k,c = kill_dummy_fix i c [] in + ast_subst (MLfix (i,fi,c)) (kill_dummy_args k 1 (MLrel 1)) with Impossible -> MLfix (i,fi,Array.map kill_dummy c)) | MLapp (MLfix (i,fi,c),a) -> let a = List.map kill_dummy a in (* Heuristics: if some arguments are implicit args, we try to eliminate the corresponding arguments of the fixpoint *) (try - let k,c = kill_dummy_fix i c (sign_of_args a) in - let fake = MLapp (MLrel 1, List.map (ast_lift 1) a) in - let fake' = kill_dummy_args k 1 fake in - ast_subst (MLfix (i,fi,c)) fake' + let k,c = kill_dummy_fix i c (sign_of_args a) in + let fake = MLapp (MLrel 1, List.map (ast_lift 1) a) in + let fake' = kill_dummy_args k 1 fake in + ast_subst (MLfix (i,fi,c)) fake' with Impossible -> MLapp(MLfix(i,fi,Array.map kill_dummy c),a)) | MLletin(id, MLfix (i,fi,c),e) -> (try - let k,c = kill_dummy_fix i c [] in - let e = kill_dummy (kill_dummy_args k 1 e) in - MLletin(id, MLfix(i,fi,c),e) + let k,c = kill_dummy_fix i c [] in + let e = kill_dummy (kill_dummy_args k 1 e) in + MLletin(id, MLfix(i,fi,c),e) with Impossible -> - MLletin(id, MLfix(i,fi,Array.map kill_dummy c),kill_dummy e)) + MLletin(id, MLfix(i,fi,Array.map kill_dummy c),kill_dummy e)) | MLletin(id,c,e) -> (try - let k,c = kill_dummy_lams [] (kill_dummy_hd c) in - let e = kill_dummy (kill_dummy_args k 1 e) in - let c = kill_dummy c in - if is_atomic c then ast_subst c e else MLletin (id, c, e) + let k,c = kill_dummy_lams [] (kill_dummy_hd c) in + let e = kill_dummy (kill_dummy_args k 1 e) in + let c = kill_dummy c in + if is_atomic c then ast_subst c e else MLletin (id, c, e) with Impossible -> MLletin(id,kill_dummy c,kill_dummy e)) | a -> ast_map kill_dummy a @@ -1329,10 +1329,10 @@ and kill_dummy_hd = function | MLlam(id,e) -> MLlam(id, kill_dummy_hd e) | MLletin(id,c,e) -> (try - let k,c = kill_dummy_lams [] (kill_dummy_hd c) in - let e = kill_dummy_hd (kill_dummy_args k 1 e) in - let c = kill_dummy c in - if is_atomic c then ast_subst c e else MLletin (id, c, e) + let k,c = kill_dummy_lams [] (kill_dummy_hd c) in + let e = kill_dummy_hd (kill_dummy_args k 1 e) in + let c = kill_dummy c in + if is_atomic c then ast_subst c e else MLletin (id, c, e) with Impossible -> MLletin(id,kill_dummy c,kill_dummy_hd e)) | a -> a @@ -1361,7 +1361,7 @@ let general_optimize_fix f ids n args m c = for i=0 to (n-1) do v.(i)<-i done; let aux i = function | MLrel j when v.(j-1)>=0 -> - if ast_occurs (j+1) c then raise Impossible else v.(j-1)<-(-i-1) + if ast_occurs (j+1) c then raise Impossible else v.(j-1)<-(-i-1) | _ -> raise Impossible in List.iteri aux args; let args_f = List.rev_map (fun i -> MLrel (i+m+1)) (Array.to_list v) in @@ -1377,19 +1377,19 @@ let optimize_fix a = if Int.equal n 0 then a else match a' with | MLfix(_,[|f|],[|c|]) -> - let new_f = MLapp (MLrel (n+1),eta_args n) in - let new_c = named_lams ids (normalize (ast_subst new_f c)) - in MLfix(0,[|f|],[|new_c|]) + let new_f = MLapp (MLrel (n+1),eta_args n) in + let new_c = named_lams ids (normalize (ast_subst new_f c)) + in MLfix(0,[|f|],[|new_c|]) | MLapp(a',args) -> - let m = List.length args in - (match a' with - | MLfix(_,_,_) when - (test_eta_args_lift 0 n args) && not (ast_occurs_itvl 1 m a') - -> a' - | MLfix(_,[|f|],[|c|]) -> - (try general_optimize_fix f ids n args m c - with Impossible -> a) - | _ -> a) + let m = List.length args in + (match a' with + | MLfix(_,_,_) when + (test_eta_args_lift 0 n args) && not (ast_occurs_itvl 1 m a') + -> a' + | MLfix(_,[|f|],[|c|]) -> + (try general_optimize_fix f ids n args m c + with Impossible -> a) + | _ -> a) | _ -> a (*S Inlining. *) @@ -1463,12 +1463,12 @@ let rec non_stricts add cand = function (* so we make an union (in fact a merge). *) let cand = non_stricts false cand t in Array.fold_left - (fun c (i,_,t)-> - let n = List.length i in - let cand = lift n cand in - let cand = pop n (non_stricts add cand t) in - List.merge Int.compare cand c) [] v - (* [merge] may duplicates some indices, but I don't mind. *) + (fun c (i,_,t)-> + let n = List.length i in + let cand = lift n cand in + let cand = pop n (non_stricts add cand t) in + List.merge Int.compare cand c) [] v + (* [merge] may duplicates some indices, but I don't mind. *) | MLmagic t -> non_stricts add cand t | _ -> diff --git a/plugins/extraction/modutil.ml b/plugins/extraction/modutil.ml index fe49bfc1ec..ec39beb03b 100644 --- a/plugins/extraction/modutil.ml +++ b/plugins/extraction/modutil.ml @@ -36,19 +36,19 @@ let se_iter do_decl do_spec do_mp = | MTident mp -> do_mp mp | MTfunsig (_,mt,mt') -> mt_iter mt; mt_iter mt' | MTwith (mt,ML_With_type(idl,l,t))-> - let mp_mt = msid_of_mt mt in - let l',idl' = List.sep_last idl in - let mp_w = - List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl' - in + let mp_mt = msid_of_mt mt in + let l',idl' = List.sep_last idl in + let mp_w = + List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl' + in let r = GlobRef.ConstRef (Constant.make2 mp_w (Label.of_id l')) in mt_iter mt; do_spec (Stype(r,l,Some t)) | MTwith (mt,ML_With_module(idl,mp))-> let mp_mt = msid_of_mt mt in let mp_w = - List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl - in - mt_iter mt; do_mp mp_w; do_mp mp + List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl + in + mt_iter mt; do_mp mp_w; do_mp mp | MTsig (_, sign) -> List.iter spec_iter sign and spec_iter = function | (_,Spec s) -> do_spec s @@ -58,7 +58,7 @@ let se_iter do_decl do_spec do_mp = let rec se_iter = function | (_,SEdecl d) -> do_decl d | (_,SEmodule m) -> - me_iter m.ml_mod_expr; mt_iter m.ml_mod_type + me_iter m.ml_mod_expr; mt_iter m.ml_mod_type | (_,SEmodtype m) -> mt_iter m and me_iter = function | MEident mp -> do_mp mp @@ -103,8 +103,8 @@ let ast_iter_references do_term do_cons do_type a = | MLglob r -> do_term r | MLcons (_,r,_) -> do_cons r | MLcase (ty,_,v) -> - type_iter_references do_type ty; - Array.iter (fun (_,p,_) -> patt_iter_references do_cons p) v + type_iter_references do_type ty; + Array.iter (fun (_,p,_) -> patt_iter_references do_cons p) v | MLrel _ | MLlam _ | MLapp _ | MLletin _ | MLtuple _ | MLfix _ | MLexn _ | MLdummy _ | MLaxiom | MLmagic _ | MLuint _ | MLfloat _ -> () @@ -118,7 +118,7 @@ let ind_iter_references do_term do_cons do_type kn ind = if lang () == Ocaml then (match ind.ind_equiv with | Miniml.Equiv kne -> do_type (GlobRef.IndRef (MutInd.make1 kne, snd ip)); - | _ -> ()); + | _ -> ()); Array.iteri (fun j -> cons_iter (ip,j+1)) p.ip_types in if lang () == Ocaml then record_iter_references do_term ind.ind_kind; @@ -132,7 +132,7 @@ let decl_iter_references do_term do_cons do_type = | Dtype (r,_,t) -> do_type r; type_iter t | Dterm (r,a,t) -> do_term r; ast_iter a; type_iter t | Dfix(rv,c,t) -> - Array.iter do_term rv; Array.iter ast_iter c; Array.iter type_iter t + Array.iter do_term rv; Array.iter ast_iter c; Array.iter type_iter t let spec_iter_references do_term do_cons do_type = function | Sind (kn,ind) -> ind_iter_references do_term do_cons do_type kn ind @@ -163,7 +163,7 @@ let rec type_search f = function let decl_type_search f = function | Dind (_,{ind_packets=p}) -> Array.iter - (fun {ip_types=v} -> Array.iter (List.iter (type_search f)) v) p + (fun {ip_types=v} -> Array.iter (List.iter (type_search f)) v) p | Dterm (_,_,u) -> type_search f u | Dfix (_,_,v) -> Array.iter (type_search f) v | Dtype (_,_,u) -> type_search f u @@ -171,7 +171,7 @@ let decl_type_search f = function let spec_type_search f = function | Sind (_,{ind_packets=p}) -> Array.iter - (fun {ip_types=v} -> Array.iter (List.iter (type_search f)) v) p + (fun {ip_types=v} -> Array.iter (List.iter (type_search f)) v) p | Stype (_,_,ot) -> Option.iter (type_search f) ot | Sval (_,u) -> type_search f u @@ -195,7 +195,7 @@ let rec msig_of_ms = function | (l,SEdecl (Dfix (rv,_,tv))) :: ms -> let msig = ref (msig_of_ms ms) in for i = Array.length rv - 1 downto 0 do - msig := (l,Spec (Sval (rv.(i),tv.(i))))::!msig + msig := (l,Spec (Sval (rv.(i),tv.(i))))::!msig done; !msig | (l,SEmodule m) :: ms -> (l,Smodule m.ml_mod_type) :: (msig_of_ms ms) @@ -229,13 +229,13 @@ let get_decl_in_structure r struc = let rec go ll sel = match ll with | [] -> assert false | l :: ll -> - match search_structure l (not (List.is_empty ll)) sel with - | SEdecl d -> d - | SEmodtype m -> assert false - | SEmodule m -> - match m.ml_mod_expr with - | MEstruct (_,sel) -> go ll sel - | _ -> error_not_visible r + match search_structure l (not (List.is_empty ll)) sel with + | SEdecl d -> d + | SEmodtype m -> assert false + | SEmodule m -> + match m.ml_mod_expr with + | MEstruct (_,sel) -> go ll sel + | _ -> error_not_visible r in go ll sel with Not_found -> anomaly (Pp.str "reference not found in extracted structure.") @@ -258,7 +258,7 @@ let dfix_to_mlfix rv av i = in let rec subst n t = match t with | MLglob ((GlobRef.ConstRef kn) as refe) -> - (try MLrel (n + (Refmap'.find refe s)) with Not_found -> t) + (try MLrel (n + (Refmap'.find refe s)) with Not_found -> t) | _ -> ast_map_lift subst n t in let ids = Array.map (fun r -> Label.to_id (label_of_r r)) rv in @@ -277,9 +277,9 @@ let rec optim_se top to_appear s = function let i = inline r a in if i then s := Refmap'.add r a !s; let d = match dump_unused_vars (optimize_fix a) with - | MLfix (0, _, [|c|]) -> - Dfix ([|r|], [|ast_subst (MLglob r) c|], [|t|]) - | a -> Dterm (r, a, t) + | MLfix (0, _, [|c|]) -> + Dfix ([|r|], [|ast_subst (MLglob r) c|], [|t|]) + | a -> Dterm (r, a, t) in (l,SEdecl d) :: (optim_se top to_appear s lse) | (l,SEdecl (Dfix (rv,av,tv))) :: lse -> @@ -287,8 +287,8 @@ let rec optim_se top to_appear s = function (* This fake body ensures that no fixpoint will be auto-inlined. *) let fake_body = MLfix (0,[||],[||]) in for i = 0 to Array.length rv - 1 do - if inline rv.(i) fake_body - then s := Refmap'.add rv.(i) (dfix_to_mlfix rv av i) !s + if inline rv.(i) fake_body + then s := Refmap'.add rv.(i) (dfix_to_mlfix rv av i) !s done; let av' = Array.map dump_unused_vars av in (l,SEdecl (Dfix (rv, av', tv))) :: (optim_se top to_appear s lse) @@ -343,7 +343,7 @@ let compute_deps_decl = function | Dterm (r,u,t) -> type_iter_references add_needed t; if not (is_custom r) then - ast_iter_references add_needed add_needed add_needed u + ast_iter_references add_needed add_needed add_needed u | Dfix _ as d -> decl_iter_references add_needed add_needed add_needed d @@ -370,10 +370,10 @@ let rec depcheck_se = function List.iter found_needed refs'; (* Hack to avoid extracting unused part of a Dfix *) match d with - | Dfix (rv,trms,tys) when (List.for_all is_custom refs') -> - let trms' = Array.make (Array.length rv) (MLexn "UNUSED") in - ((l,SEdecl (Dfix (rv,trms',tys))) :: se') - | _ -> (compute_deps_decl d; t::se') + | Dfix (rv,trms,tys) when (List.for_all is_custom refs') -> + let trms' = Array.make (Array.length rv) (MLexn "UNUSED") in + ((l,SEdecl (Dfix (rv,trms',tys))) :: se') + | _ -> (compute_deps_decl d; t::se') end | t :: se -> let se' = depcheck_se se in diff --git a/plugins/extraction/ocaml.ml b/plugins/extraction/ocaml.ml index 34ddf57b40..66429833b9 100644 --- a/plugins/extraction/ocaml.ml +++ b/plugins/extraction/ocaml.ml @@ -163,16 +163,16 @@ let pp_type par vl t = | Tvar i -> (try pp_tvar (List.nth vl (pred i)) with Failure _ -> (str "'a" ++ int i)) | Tglob (r,[a1;a2]) when is_infix r -> - pp_par par (pp_rec true a1 ++ str (get_infix r) ++ pp_rec true a2) + pp_par par (pp_rec true a1 ++ str (get_infix r) ++ pp_rec true a2) | Tglob (r,[]) -> pp_global Type r | Tglob (GlobRef.IndRef(kn,0),l) - when not (keep_singleton ()) && MutInd.equal kn (mk_ind "Coq.Init.Specif" "sig") -> - pp_tuple_light pp_rec l + when not (keep_singleton ()) && MutInd.equal kn (mk_ind "Coq.Init.Specif" "sig") -> + pp_tuple_light pp_rec l | Tglob (r,l) -> - pp_tuple_light pp_rec l ++ spc () ++ pp_global Type r + pp_tuple_light pp_rec l ++ spc () ++ pp_global Type r | Tarr (t1,t2) -> - pp_par par - (pp_rec true t1 ++ spc () ++ str "->" ++ spc () ++ pp_rec false t2) + pp_par par + (pp_rec true t1 ++ spc () ++ str "->" ++ spc () ++ pp_rec false t2) | Tdummy _ -> str "__" | Tunknown -> str "__" in @@ -209,101 +209,101 @@ let rec pp_expr par env args = and apply2 st = pp_apply2 st par args in function | MLrel n -> - let id = get_db_name n env in + let id = get_db_name n env in (* Try to survive to the occurrence of a Dummy rel. TODO: we should get rid of this hack (cf. #592) *) let id = if Id.equal id dummy_name then Id.of_string "__" else id in apply (Id.print id) | MLapp (f,args') -> - let stl = List.map (pp_expr true env []) args' in + let stl = List.map (pp_expr true env []) args' in pp_expr par env (stl @ args) f | MLlam _ as a -> - let fl,a' = collect_lams a in - let fl = List.map id_of_mlid fl in - let fl,env' = push_vars fl env in - let st = pp_abst (List.rev fl) ++ pp_expr false env' [] a' in - apply2 st + let fl,a' = collect_lams a in + let fl = List.map id_of_mlid fl in + let fl,env' = push_vars fl env in + let st = pp_abst (List.rev fl) ++ pp_expr false env' [] a' in + apply2 st | MLletin (id,a1,a2) -> - let i,env' = push_vars [id_of_mlid id] env in - let pp_id = Id.print (List.hd i) - and pp_a1 = pp_expr false env [] a1 - and pp_a2 = pp_expr (not par && expr_needs_par a2) env' [] a2 in - hv 0 (apply2 (pp_letin pp_id pp_a1 pp_a2)) + let i,env' = push_vars [id_of_mlid id] env in + let pp_id = Id.print (List.hd i) + and pp_a1 = pp_expr false env [] a1 + and pp_a2 = pp_expr (not par && expr_needs_par a2) env' [] a2 in + hv 0 (apply2 (pp_letin pp_id pp_a1 pp_a2)) | MLglob r -> apply (pp_global Term r) | MLfix (i,ids,defs) -> - let ids',env' = push_vars (List.rev (Array.to_list ids)) env in - pp_fix par env' i (Array.of_list (List.rev ids'),defs) args + let ids',env' = push_vars (List.rev (Array.to_list ids)) env in + pp_fix par env' i (Array.of_list (List.rev ids'),defs) args | MLexn s -> - (* An [MLexn] may be applied, but I don't really care. *) - pp_par par (str "assert false" ++ spc () ++ str ("(* "^s^" *)")) + (* An [MLexn] may be applied, but I don't really care. *) + pp_par par (str "assert false" ++ spc () ++ str ("(* "^s^" *)")) | MLdummy k -> (* An [MLdummy] may be applied, but I don't really care. *) (match msg_of_implicit k with | "" -> str "__" | s -> str "__" ++ spc () ++ str ("(* "^s^" *)")) | MLmagic a -> - pp_apply (str "Obj.magic") par (pp_expr true env [] a :: args) + pp_apply (str "Obj.magic") par (pp_expr true env [] a :: args) | MLaxiom -> - pp_par par (str "failwith \"AXIOM TO BE REALIZED\"") + pp_par par (str "failwith \"AXIOM TO BE REALIZED\"") | MLcons (_,r,a) as c -> assert (List.is_empty args); begin match a with - | _ when is_native_char c -> pp_native_char c - | [a1;a2] when is_infix r -> - let pp = pp_expr true env [] in - pp_par par (pp a1 ++ str (get_infix r) ++ pp a2) - | _ when is_coinductive r -> - let ne = not (List.is_empty a) in - let tuple = space_if ne ++ pp_tuple (pp_expr true env []) a in - pp_par par (str "lazy " ++ pp_par ne (pp_global Cons r ++ tuple)) - | [] -> pp_global Cons r - | _ -> - let fds = get_record_fields r in - if not (List.is_empty fds) then - pp_record_pat (pp_fields r fds, List.map (pp_expr true env []) a) - else - let tuple = pp_tuple (pp_expr true env []) a in - if String.is_empty (str_global Cons r) (* hack Extract Inductive prod *) - then tuple - else pp_par par (pp_global Cons r ++ spc () ++ tuple) - end + | _ when is_native_char c -> pp_native_char c + | [a1;a2] when is_infix r -> + let pp = pp_expr true env [] in + pp_par par (pp a1 ++ str (get_infix r) ++ pp a2) + | _ when is_coinductive r -> + let ne = not (List.is_empty a) in + let tuple = space_if ne ++ pp_tuple (pp_expr true env []) a in + pp_par par (str "lazy " ++ pp_par ne (pp_global Cons r ++ tuple)) + | [] -> pp_global Cons r + | _ -> + let fds = get_record_fields r in + if not (List.is_empty fds) then + pp_record_pat (pp_fields r fds, List.map (pp_expr true env []) a) + else + let tuple = pp_tuple (pp_expr true env []) a in + if String.is_empty (str_global Cons r) (* hack Extract Inductive prod *) + then tuple + else pp_par par (pp_global Cons r ++ spc () ++ tuple) + end | MLtuple l -> assert (List.is_empty args); pp_boxed_tuple (pp_expr true env []) l | MLcase (_, t, pv) when is_custom_match pv -> if not (is_regular_match pv) then - user_err Pp.(str "Cannot mix yet user-given match and general patterns."); - let mkfun (ids,_,e) = - if not (List.is_empty ids) then named_lams (List.rev ids) e - else dummy_lams (ast_lift 1 e) 1 - in - let pp_branch tr = pp_expr true env [] (mkfun tr) ++ fnl () in - let inner = - str (find_custom_match pv) ++ fnl () ++ - prvect pp_branch pv ++ - pp_expr true env [] t - in - apply2 (hov 2 inner) + user_err Pp.(str "Cannot mix yet user-given match and general patterns."); + let mkfun (ids,_,e) = + if not (List.is_empty ids) then named_lams (List.rev ids) e + else dummy_lams (ast_lift 1 e) 1 + in + let pp_branch tr = pp_expr true env [] (mkfun tr) ++ fnl () in + let inner = + str (find_custom_match pv) ++ fnl () ++ + prvect pp_branch pv ++ + pp_expr true env [] t + in + apply2 (hov 2 inner) | MLcase (typ, t, pv) -> let head = - if not (is_coinductive_type typ) then pp_expr false env [] t - else (str "Lazy.force" ++ spc () ++ pp_expr true env [] t) - in - (* First, can this match be printed as a mere record projection ? *) + if not (is_coinductive_type typ) then pp_expr false env [] t + else (str "Lazy.force" ++ spc () ++ pp_expr true env [] t) + in + (* First, can this match be printed as a mere record projection ? *) (try pp_record_proj par env typ t pv args - with Impossible -> - (* Second, can this match be printed as a let-in ? *) - if Int.equal (Array.length pv) 1 then - let s1,s2 = pp_one_pat env pv.(0) in - hv 0 (apply2 (pp_letin s1 head s2)) - else - (* Third, can this match be printed as [if ... then ... else] ? *) - (try apply2 (pp_ifthenelse env head pv) - with Not_found -> - (* Otherwise, standard match *) - apply2 - (v 0 (str "match " ++ head ++ str " with" ++ fnl () ++ - pp_pat env pv)))) + with Impossible -> + (* Second, can this match be printed as a let-in ? *) + if Int.equal (Array.length pv) 1 then + let s1,s2 = pp_one_pat env pv.(0) in + hv 0 (apply2 (pp_letin s1 head s2)) + else + (* Third, can this match be printed as [if ... then ... else] ? *) + (try apply2 (pp_ifthenelse env head pv) + with Not_found -> + (* Otherwise, standard match *) + apply2 + (v 0 (str "match " ++ head ++ str " with" ++ fnl () ++ + pp_pat env pv)))) | MLuint i -> assert (args=[]); str "(" ++ str (Uint63.compile i) ++ str ")" @@ -381,9 +381,9 @@ and pp_ifthenelse env expr pv = match pv with -> hv 0 (hov 2 (str "if " ++ expr) ++ spc () ++ hov 2 (str "then " ++ - hov 2 (pp_expr (expr_needs_par the) env [] the)) ++ spc () ++ - hov 2 (str "else " ++ - hov 2 (pp_expr (expr_needs_par els) env [] els))) + hov 2 (pp_expr (expr_needs_par the) env [] the)) ++ spc () ++ + hov 2 (str "else " ++ + hov 2 (pp_expr (expr_needs_par els) env [] els))) | _ -> raise Not_found and pp_one_pat env (ids,p,t) = @@ -404,20 +404,20 @@ and pp_function env t = let bl,env' = push_vars (List.map id_of_mlid bl) env in match t' with | MLcase(Tglob(r,_),MLrel 1,pv) when - not (is_coinductive r) && List.is_empty (get_record_fields r) && - not (is_custom_match pv) -> - if not (ast_occurs 1 (MLcase(Tunknown,MLaxiom,pv))) then - pr_binding (List.rev (List.tl bl)) ++ - str " = function" ++ fnl () ++ - v 0 (pp_pat env' pv) - else + not (is_coinductive r) && List.is_empty (get_record_fields r) && + not (is_custom_match pv) -> + if not (ast_occurs 1 (MLcase(Tunknown,MLaxiom,pv))) then + pr_binding (List.rev (List.tl bl)) ++ + str " = function" ++ fnl () ++ + v 0 (pp_pat env' pv) + else pr_binding (List.rev bl) ++ str " = match " ++ Id.print (List.hd bl) ++ str " with" ++ fnl () ++ - v 0 (pp_pat env' pv) + v 0 (pp_pat env' pv) | _ -> pr_binding (List.rev bl) ++ - str " =" ++ fnl () ++ str " " ++ - hov 2 (pp_expr false env' [] t') + str " =" ++ fnl () ++ str " " ++ + hov 2 (pp_expr false env' [] t') (*s names of the functions ([ids]) are already pushed in [env], and passed here just for convenience. *) @@ -425,12 +425,12 @@ and pp_function env t = and pp_fix par env i (ids,bl) args = pp_par par (v 0 (str "let rec " ++ - prvect_with_sep - (fun () -> fnl () ++ str "and ") - (fun (fi,ti) -> Id.print fi ++ pp_function env ti) - (Array.map2 (fun id b -> (id,b)) ids bl) ++ - fnl () ++ - hov 2 (str "in " ++ pp_apply (Id.print ids.(i)) false args))) + prvect_with_sep + (fun () -> fnl () ++ str "and ") + (fun (fi,ti) -> Id.print fi ++ pp_function env ti) + (Array.map2 (fun id b -> (id,b)) ids bl) ++ + fnl () ++ + hov 2 (str "in " ++ pp_apply (Id.print ids.(i)) false args))) (* Ad-hoc double-newline in v boxes, with enough negative whitespace to avoid indenting the intermediate blank line *) @@ -451,19 +451,19 @@ let pp_Dfix (rv,c,t) = if i >= Array.length rv then mt () else let void = is_inline_custom rv.(i) || - (not (is_custom rv.(i)) && + (not (is_custom rv.(i)) && match c.(i) with MLexn "UNUSED" -> true | _ -> false) in if void then pp init (i+1) else - let def = - if is_custom rv.(i) then str " = " ++ str (find_custom rv.(i)) - else pp_function (empty_env ()) c.(i) - in - (if init then mt () else cut2 ()) ++ - pp_val names.(i) t.(i) ++ - str (if init then "let rec " else "and ") ++ names.(i) ++ def ++ - pp false (i+1) + let def = + if is_custom rv.(i) then str " = " ++ str (find_custom rv.(i)) + else pp_function (empty_env ()) c.(i) + in + (if init then mt () else cut2 ()) ++ + pp_val names.(i) t.(i) ++ + str (if init then "let rec " else "and ") ++ names.(i) ++ def ++ + pp false (i+1) in pp true 0 (*s Pretty-printing of inductive types declaration. *) @@ -481,9 +481,9 @@ let pp_one_ind prefix ip_equiv pl name cnames ctyps = let pp_constructor i typs = (if Int.equal i 0 then mt () else fnl ()) ++ hov 3 (str "| " ++ cnames.(i) ++ - (if List.is_empty typs then mt () else str " of ") ++ - prlist_with_sep - (fun () -> spc () ++ str "* ") (pp_type true pl) typs) + (if List.is_empty typs then mt () else str " of ") ++ + prlist_with_sep + (fun () -> spc () ++ str "* ") (pp_type true pl) typs) in pp_parameters pl ++ str prefix ++ name ++ pp_equiv pl name ip_equiv ++ str " =" ++ @@ -494,16 +494,16 @@ let pp_logical_ind packet = pp_comment (Id.print packet.ip_typename ++ str " : logical inductive") ++ fnl () ++ pp_comment (str "with constructors : " ++ - prvect_with_sep spc Id.print packet.ip_consnames) ++ + prvect_with_sep spc Id.print packet.ip_consnames) ++ fnl () let pp_singleton kn packet = let name = pp_global Type (GlobRef.IndRef (kn,0)) in let l = rename_tvars keywords packet.ip_vars in hov 2 (str "type " ++ pp_parameters l ++ name ++ str " =" ++ spc () ++ - pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++ - pp_comment (str "singleton inductive, whose constructor was " ++ - Id.print packet.ip_consnames.(0))) + pp_type false l (List.hd packet.ip_types.(0)) ++ fnl () ++ + pp_comment (str "singleton inductive, whose constructor was " ++ + Id.print packet.ip_consnames.(0))) let pp_record kn fields ip_equiv packet = let ind = GlobRef.IndRef (kn,0) in @@ -514,7 +514,7 @@ let pp_record kn fields ip_equiv packet = str "type " ++ pp_parameters pl ++ name ++ pp_equiv pl name ip_equiv ++ str " = { "++ hov 0 (prlist_with_sep (fun () -> str ";" ++ spc ()) - (fun (p,t) -> p ++ str " : " ++ pp_type true pl t) l) + (fun (p,t) -> p ++ str " : " ++ pp_type true pl t) l) ++ str " }" let pp_coind pl name = @@ -536,7 +536,7 @@ let pp_ind co kn ind = Array.mapi (fun i p -> if p.ip_logical then [||] else Array.mapi (fun j _ -> pp_global Cons (GlobRef.ConstructRef ((kn,i),j+1))) - p.ip_types) + p.ip_types) ind.ind_packets in let rec pp i kwd = @@ -548,9 +548,9 @@ let pp_ind co kn ind = if is_custom (GlobRef.IndRef ip) then pp (i+1) kwd else if p.ip_logical then pp_logical_ind p ++ pp (i+1) kwd else - kwd ++ (if co then pp_coind p.ip_vars names.(i) else mt ()) ++ - pp_one_ind prefix ip_equiv p.ip_vars names.(i) cnames.(i) p.ip_types ++ - pp (i+1) nextkwd + kwd ++ (if co then pp_coind p.ip_vars names.(i) else mt ()) ++ + pp_one_ind prefix ip_equiv p.ip_vars names.(i) cnames.(i) p.ip_types ++ + pp (i+1) nextkwd in pp 0 initkwd @@ -570,26 +570,26 @@ let pp_decl = function | Dind (kn,i) -> pp_mind kn i | Dtype (r, l, t) -> let name = pp_global Type r in - let l = rename_tvars keywords l in + let l = rename_tvars keywords l in let ids, def = - try - let ids,s = find_type_custom r in - pp_string_parameters ids, str " =" ++ spc () ++ str s - with Not_found -> - pp_parameters l, - if t == Taxiom then str " (* AXIOM TO BE REALIZED *)" - else str " =" ++ spc () ++ pp_type false l t - in - hov 2 (str "type " ++ ids ++ name ++ def) + try + let ids,s = find_type_custom r in + pp_string_parameters ids, str " =" ++ spc () ++ str s + with Not_found -> + pp_parameters l, + if t == Taxiom then str " (* AXIOM TO BE REALIZED *)" + else str " =" ++ spc () ++ pp_type false l t + in + hov 2 (str "type " ++ ids ++ name ++ def) | Dterm (r, a, t) -> - let def = - if is_custom r then str (" = " ^ find_custom r) - else pp_function (empty_env ()) a - in - let name = pp_global Term r in + let def = + if is_custom r then str (" = " ^ find_custom r) + else pp_function (empty_env ()) a + in + let name = pp_global Term r in pp_val name t ++ hov 0 (str "let " ++ name ++ def ++ mt ()) | Dfix (rv,defs,typs) -> - pp_Dfix (rv,defs,typs) + pp_Dfix (rv,defs,typs) let pp_spec = function | Sval (r,_) when is_inline_custom r -> mt () @@ -603,15 +603,15 @@ let pp_spec = function let name = pp_global Type r in let l = rename_tvars keywords vl in let ids, def = - try - let ids, s = find_type_custom r in - pp_string_parameters ids, str " =" ++ spc () ++ str s - with Not_found -> - let ids = pp_parameters l in - match ot with - | None -> ids, mt () - | Some Taxiom -> ids, str " (* AXIOM TO BE REALIZED *)" - | Some t -> ids, str " =" ++ spc () ++ pp_type false l t + try + let ids, s = find_type_custom r in + pp_string_parameters ids, str " =" ++ spc () ++ str s + with Not_found -> + let ids = pp_parameters l in + match ot with + | None -> ids, mt () + | Some Taxiom -> ids, str " (* AXIOM TO BE REALIZED *)" + | Some t -> ids, str " =" ++ spc () ++ pp_type false l t in hov 2 (str "type " ++ ids ++ name ++ def) @@ -621,8 +621,8 @@ let rec pp_specif = function (match Common.get_duplicate (top_visible_mp ()) l with | None -> pp_spec s | Some ren -> - hov 1 (str ("module "^ren^" : sig") ++ fnl () ++ pp_spec s) ++ - fnl () ++ str "end" ++ fnl () ++ + hov 1 (str ("module "^ren^" : sig") ++ fnl () ++ pp_spec s) ++ + fnl () ++ str "end" ++ fnl () ++ str ("include module type of struct include "^ren^" end")) | (l,Smodule mt) -> let def = pp_module_type [] mt in @@ -670,7 +670,7 @@ and pp_module_type params = function let mp_mt = msid_of_mt mt in let l,idl' = List.sep_last idl in let mp_w = - List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl' + List.fold_left (fun mp l -> MPdot(mp,Label.of_id l)) mp_mt idl' in let r = GlobRef.ConstRef (Constant.make2 mp_w (Label.of_id l)) in push_visible mp_mt []; @@ -680,7 +680,7 @@ and pp_module_type params = function | MTwith(mt,ML_With_module(idl,mp)) -> let mp_mt = msid_of_mt mt in let mp_w = - List.fold_left (fun mp id -> MPdot(mp,Label.of_id id)) mp_mt idl + List.fold_left (fun mp id -> MPdot(mp,Label.of_id id)) mp_mt idl in push_visible mp_mt []; let pp_w = str " with module " ++ pp_modname mp_w in @@ -694,20 +694,20 @@ let rec pp_structure_elem = function (match Common.get_duplicate (top_visible_mp ()) l with | None -> pp_decl d | Some ren -> - hov 1 (str ("module "^ren^" = struct") ++ fnl () ++ pp_decl d) ++ - fnl () ++ str "end" ++ fnl () ++ str ("include "^ren)) + hov 1 (str ("module "^ren^" = struct") ++ fnl () ++ pp_decl d) ++ + fnl () ++ str "end" ++ fnl () ++ str ("include "^ren)) | (l,SEmodule m) -> let typ = (* virtual printing of the type, in order to have a correct mli later*) - if Common.get_phase () == Pre then - str ": " ++ pp_module_type [] m.ml_mod_type - else mt () + if Common.get_phase () == Pre then + str ": " ++ pp_module_type [] m.ml_mod_type + else mt () in let def = pp_module_expr [] m.ml_mod_expr in let name = pp_modname (MPdot (top_visible_mp (), l)) in hov 1 - (str "module " ++ name ++ typ ++ str " =" ++ - (if is_short m.ml_mod_expr then spc () else fnl ()) ++ def) ++ + (str "module " ++ name ++ typ ++ str " =" ++ + (if is_short m.ml_mod_expr then spc () else fnl ()) ++ def) ++ (match Common.get_duplicate (top_visible_mp ()) l with | Some ren -> fnl () ++ str ("module "^ren^" = ") ++ name | None -> mt ()) diff --git a/plugins/extraction/scheme.ml b/plugins/extraction/scheme.ml index c341ec8d57..c41b0d7a02 100644 --- a/plugins/extraction/scheme.ml +++ b/plugins/extraction/scheme.ml @@ -50,13 +50,13 @@ let pp_abst st = function | [] -> assert false | [id] -> paren (str "lambda " ++ paren (pr_id id) ++ spc () ++ st) | l -> paren - (str "lambdas " ++ paren (prlist_with_sep spc pr_id l) ++ spc () ++ st) + (str "lambdas " ++ paren (prlist_with_sep spc pr_id l) ++ spc () ++ st) let pp_apply st _ = function | [] -> st | [a] -> hov 2 (paren (st ++ spc () ++ a)) | args -> hov 2 (paren (str "@ " ++ st ++ - (prlist_strict (fun x -> spc () ++ x) args))) + (prlist_strict (fun x -> spc () ++ x) args))) (*s The pretty-printer for Scheme syntax *) @@ -68,66 +68,66 @@ let rec pp_expr env args = let apply st = pp_apply st true args in function | MLrel n -> - let id = get_db_name n env in apply (pr_id id) + let id = get_db_name n env in apply (pr_id id) | MLapp (f,args') -> - let stl = List.map (pp_expr env []) args' in + let stl = List.map (pp_expr env []) args' in pp_expr env (stl @ args) f | MLlam _ as a -> - let fl,a' = collect_lams a in - let fl,env' = push_vars (List.map id_of_mlid fl) env in - apply (pp_abst (pp_expr env' [] a') (List.rev fl)) + let fl,a' = collect_lams a in + let fl,env' = push_vars (List.map id_of_mlid fl) env in + apply (pp_abst (pp_expr env' [] a') (List.rev fl)) | MLletin (id,a1,a2) -> - let i,env' = push_vars [id_of_mlid id] env in - apply - (hv 0 - (hov 2 - (paren - (str "let " ++ - paren - (paren - (pr_id (List.hd i) ++ spc () ++ pp_expr env [] a1)) - ++ spc () ++ hov 0 (pp_expr env' [] a2))))) + let i,env' = push_vars [id_of_mlid id] env in + apply + (hv 0 + (hov 2 + (paren + (str "let " ++ + paren + (paren + (pr_id (List.hd i) ++ spc () ++ pp_expr env [] a1)) + ++ spc () ++ hov 0 (pp_expr env' [] a2))))) | MLglob r -> - apply (pp_global Term r) + apply (pp_global Term r) | MLcons (_,r,args') -> - assert (List.is_empty args); - let st = - str "`" ++ - paren (pp_global Cons r ++ - (if List.is_empty args' then mt () else spc ()) ++ - prlist_with_sep spc (pp_cons_args env) args') - in - if is_coinductive r then paren (str "delay " ++ st) else st + assert (List.is_empty args); + let st = + str "`" ++ + paren (pp_global Cons r ++ + (if List.is_empty args' then mt () else spc ()) ++ + prlist_with_sep spc (pp_cons_args env) args') + in + if is_coinductive r then paren (str "delay " ++ st) else st | MLtuple _ -> user_err Pp.(str "Cannot handle tuples in Scheme yet.") | MLcase (_,_,pv) when not (is_regular_match pv) -> - user_err Pp.(str "Cannot handle general patterns in Scheme yet.") + user_err Pp.(str "Cannot handle general patterns in Scheme yet.") | MLcase (_,t,pv) when is_custom_match pv -> - let mkfun (ids,_,e) = - if not (List.is_empty ids) then named_lams (List.rev ids) e - else dummy_lams (ast_lift 1 e) 1 - in - apply - (paren - (hov 2 - (str (find_custom_match pv) ++ fnl () ++ - prvect (fun tr -> pp_expr env [] (mkfun tr) ++ fnl ()) pv - ++ pp_expr env [] t))) + let mkfun (ids,_,e) = + if not (List.is_empty ids) then named_lams (List.rev ids) e + else dummy_lams (ast_lift 1 e) 1 + in + apply + (paren + (hov 2 + (str (find_custom_match pv) ++ fnl () ++ + prvect (fun tr -> pp_expr env [] (mkfun tr) ++ fnl ()) pv + ++ pp_expr env [] t))) | MLcase (typ,t, pv) -> let e = - if not (is_coinductive_type typ) then pp_expr env [] t - else paren (str "force" ++ spc () ++ pp_expr env [] t) - in - apply (v 3 (paren (str "match " ++ e ++ fnl () ++ pp_pat env pv))) + if not (is_coinductive_type typ) then pp_expr env [] t + else paren (str "force" ++ spc () ++ pp_expr env [] t) + in + apply (v 3 (paren (str "match " ++ e ++ fnl () ++ pp_pat env pv))) | MLfix (i,ids,defs) -> - let ids',env' = push_vars (List.rev (Array.to_list ids)) env in - pp_fix env' i (Array.of_list (List.rev ids'),defs) args + let ids',env' = push_vars (List.rev (Array.to_list ids)) env in + pp_fix env' i (Array.of_list (List.rev ids'),defs) args | MLexn s -> - (* An [MLexn] may be applied, but I don't really care. *) - paren (str "error" ++ spc () ++ qs s) + (* An [MLexn] may be applied, but I don't really care. *) + paren (str "error" ++ spc () ++ qs s) | MLdummy _ -> - str "__" (* An [MLdummy] may be applied, but I don't really care. *) + str "__" (* An [MLdummy] may be applied, but I don't really care. *) | MLmagic a -> - pp_expr env args a + pp_expr env args a | MLaxiom -> paren (str "error \"AXIOM TO BE REALIZED\"") | MLuint _ -> paren (str "Prelude.error \"EXTRACTION OF UINT NOT IMPLEMENTED\"") @@ -137,8 +137,8 @@ let rec pp_expr env args = and pp_cons_args env = function | MLcons (_,r,args) when is_coinductive r -> paren (pp_global Cons r ++ - (if List.is_empty args then mt () else spc ()) ++ - prlist_with_sep spc (pp_cons_args env) args) + (if List.is_empty args then mt () else spc ()) ++ + prlist_with_sep spc (pp_cons_args env) args) | e -> str "," ++ pp_expr env [] e and pp_one_pat env (ids,p,t) = @@ -166,12 +166,12 @@ and pp_fix env j (ids,bl) args = paren (str "letrec " ++ (v 0 (paren - (prvect_with_sep fnl - (fun (fi,ti) -> - paren ((pr_id fi) ++ spc () ++ (pp_expr env [] ti))) - (Array.map2 (fun id b -> (id,b)) ids bl)) ++ - fnl () ++ - hov 2 (pp_apply (pr_id (ids.(j))) true args)))) + (prvect_with_sep fnl + (fun (fi,ti) -> + paren ((pr_id fi) ++ spc () ++ (pp_expr env [] ti))) + (Array.map2 (fun id b -> (id,b)) ids bl)) ++ + fnl () ++ + hov 2 (pp_apply (pr_id (ids.(j))) true args)))) (*s Pretty-printing of a declaration. *) @@ -180,29 +180,29 @@ let pp_decl = function | Dtype _ -> mt () | Dfix (rv, defs,_) -> let names = Array.map - (fun r -> if is_inline_custom r then mt () else pp_global Term r) rv + (fun r -> if is_inline_custom r then mt () else pp_global Term r) rv in prvecti - (fun i r -> - let void = is_inline_custom r || - (not (is_custom r) && + (fun i r -> + let void = is_inline_custom r || + (not (is_custom r) && match defs.(i) with MLexn "UNUSED" -> true | _ -> false) - in - if void then mt () - else - hov 2 - (paren (str "define " ++ names.(i) ++ spc () ++ - (if is_custom r then str (find_custom r) - else pp_expr (empty_env ()) [] defs.(i))) - ++ fnl ()) ++ fnl ()) - rv + in + if void then mt () + else + hov 2 + (paren (str "define " ++ names.(i) ++ spc () ++ + (if is_custom r then str (find_custom r) + else pp_expr (empty_env ()) [] defs.(i))) + ++ fnl ()) ++ fnl ()) + rv | Dterm (r, a, _) -> if is_inline_custom r then mt () else - hov 2 (paren (str "define " ++ pp_global Term r ++ spc () ++ - (if is_custom r then str (find_custom r) - else pp_expr (empty_env ()) [] a))) - ++ fnl2 () + hov 2 (paren (str "define " ++ pp_global Term r ++ spc () ++ + (if is_custom r then str (find_custom r) + else pp_expr (empty_env ()) [] a))) + ++ fnl2 () let rec pp_structure_elem = function | (l,SEdecl d) -> pp_decl d diff --git a/plugins/extraction/table.ml b/plugins/extraction/table.ml index be9259861a..7b64706138 100644 --- a/plugins/extraction/table.ml +++ b/plugins/extraction/table.ml @@ -289,7 +289,7 @@ let safe_pr_long_global r = with Not_found -> match r with | GlobRef.ConstRef kn -> let mp,l = Constant.repr2 kn in - str ((ModPath.to_string mp)^"."^(Label.to_string l)) + str ((ModPath.to_string mp)^"."^(Label.to_string l)) | _ -> assert false let pr_long_mp mp = @@ -339,10 +339,10 @@ let warn_extraction_opaque_accessed = let warn_extraction_opaque_as_axiom = CWarnings.create ~name:"extraction-opaque-as-axiom" ~category:"extraction" (fun lst -> strbrk "The extraction now honors the opacity constraints by default, " ++ - strbrk "the following opaque constants have been extracted as axioms :" ++ - lst ++ str "." ++ fnl () ++ - strbrk "If necessary, use \"Set Extraction AccessOpaque\" to change this." - ++ fnl ()) + strbrk "the following opaque constants have been extracted as axioms :" ++ + lst ++ str "." ++ fnl () ++ + strbrk "If necessary, use \"Set Extraction AccessOpaque\" to change this." + ++ fnl ()) let warning_opaques accessed = let opaques = Refset'.elements !opaques in @@ -375,14 +375,14 @@ let warn_extraction_inside_module = let check_inside_module () = if Lib.is_modtype () then err (str "You can't do that within a Module Type." ++ fnl () ++ - str "Close it and try again.") + str "Close it and try again.") else if Lib.is_module () then warn_extraction_inside_module () let check_inside_section () = if Global.sections_are_opened () then err (str "You can't do that within a section." ++ fnl () ++ - str "Close it and try again.") + str "Close it and try again.") let warn_extraction_reserved_identifier = CWarnings.create ~name:"extraction-reserved-identifier" ~category:"extraction" @@ -441,9 +441,9 @@ let error_MPfile_as_mod mp b = let s1 = if b then "asked" else "required" in let s2 = if b then "extract some objects of this module or\n" else "" in err (str ("Extraction of file "^(raw_string_of_modfile mp)^ - ".v as a module is "^s1^".\n"^ - "Monolithic Extraction cannot deal with this situation.\n"^ - "Please "^s2^"use (Recursive) Extraction Library instead.\n")) + ".v as a module is "^s1^".\n"^ + "Monolithic Extraction cannot deal with this situation.\n"^ + "Please "^s2^"use (Recursive) Extraction Library instead.\n")) let argnames_of_global r = let env = Global.env () in @@ -481,10 +481,10 @@ let warning_remaining_implicit k = let check_loaded_modfile mp = match base_mp mp with | MPfile dp -> if not (Library.library_is_loaded dp) then begin - match base_mp (Lib.current_mp ()) with - | MPfile dp' when not (DirPath.equal dp dp') -> + match base_mp (Lib.current_mp ()) with + | MPfile dp' when not (DirPath.equal dp dp') -> err (str "Please load library " ++ DirPath.print dp ++ str " first.") - | _ -> () + | _ -> () end | _ -> () @@ -574,11 +574,11 @@ let chg_flag n = int_flag_ref := n; opt_flag_ref := flag_of_int n let optims () = !opt_flag_ref let () = declare_bool_option - {optdepr = false; - optname = "Extraction Optimize"; - optkey = ["Extraction"; "Optimize"]; - optread = (fun () -> not (Int.equal !int_flag_ref 0)); - optwrite = (fun b -> chg_flag (if b then int_flag_init else 0))} + {optdepr = false; + optname = "Extraction Optimize"; + optkey = ["Extraction"; "Optimize"]; + optread = (fun () -> not (Int.equal !int_flag_ref 0)); + optwrite = (fun b -> chg_flag (if b then int_flag_init else 0))} let () = declare_int_option { optdepr = false; @@ -671,11 +671,11 @@ let print_extraction_inline () = (str "Extraction Inline:" ++ fnl () ++ Refset'.fold (fun r p -> - (p ++ str " " ++ safe_pr_long_global r ++ fnl ())) i' (mt ()) ++ + (p ++ str " " ++ safe_pr_long_global r ++ fnl ())) i' (mt ()) ++ str "Extraction NoInline:" ++ fnl () ++ Refset'.fold (fun r p -> - (p ++ str " " ++ safe_pr_long_global r ++ fnl ())) n (mt ())) + (p ++ str " " ++ safe_pr_long_global r ++ fnl ())) n (mt ())) (* Reset part *) @@ -708,16 +708,16 @@ let add_implicits r l = let n = List.length names in let add_arg s = function | ArgInt i -> - if 1 <= i && i <= n then Int.Set.add i s - else err (int i ++ str " is not a valid argument number for " ++ - safe_pr_global r) + if 1 <= i && i <= n then Int.Set.add i s + else err (int i ++ str " is not a valid argument number for " ++ + safe_pr_global r) | ArgId id -> try let i = List.index Name.equal (Name id) names in Int.Set.add i s with Not_found -> - err (str "No argument " ++ Id.print id ++ str " for " ++ - safe_pr_global r) + err (str "No argument " ++ Id.print id ++ str " for " ++ + safe_pr_global r) in let ints = List.fold_left add_arg Int.Set.empty l in implicits_table := Refmap'.add r ints !implicits_table @@ -854,16 +854,16 @@ let extract_constant_inline inline r ids s = let g = Smartlocate.global_with_alias r in match g with | GlobRef.ConstRef kn -> - let env = Global.env () in + let env = Global.env () in let typ, _ = Typeops.type_of_global_in_context env (GlobRef.ConstRef kn) in - let typ = Reduction.whd_all env typ in - if Reduction.is_arity env typ - then begin - let nargs = Hook.get use_type_scheme_nb_args env typ in - if not (Int.equal (List.length ids) nargs) then error_axiom_scheme g nargs - end; - Lib.add_anonymous_leaf (inline_extraction (inline,[g])); - Lib.add_anonymous_leaf (in_customs (g,ids,s)) + let typ = Reduction.whd_all env typ in + if Reduction.is_arity env typ + then begin + let nargs = Hook.get use_type_scheme_nb_args env typ in + if not (Int.equal (List.length ids) nargs) then error_axiom_scheme g nargs + end; + Lib.add_anonymous_leaf (inline_extraction (inline,[g])); + Lib.add_anonymous_leaf (in_customs (g,ids,s)) | _ -> error_constant g @@ -873,18 +873,18 @@ let extract_inductive r s l optstr = Dumpglob.add_glob ?loc:r.CAst.loc g; match g with | GlobRef.IndRef ((kn,i) as ip) -> - let mib = Global.lookup_mind kn in - let n = Array.length mib.mind_packets.(i).mind_consnames in - if not (Int.equal n (List.length l)) then error_nb_cons (); - Lib.add_anonymous_leaf (inline_extraction (true,[g])); - Lib.add_anonymous_leaf (in_customs (g,[],s)); - Option.iter (fun s -> Lib.add_anonymous_leaf (in_custom_matchs (g,s))) - optstr; - List.iteri - (fun j s -> + let mib = Global.lookup_mind kn in + let n = Array.length mib.mind_packets.(i).mind_consnames in + if not (Int.equal n (List.length l)) then error_nb_cons (); + Lib.add_anonymous_leaf (inline_extraction (true,[g])); + Lib.add_anonymous_leaf (in_customs (g,[],s)); + Option.iter (fun s -> Lib.add_anonymous_leaf (in_custom_matchs (g,s))) + optstr; + List.iteri + (fun j s -> let g = GlobRef.ConstructRef (ip,succ j) in - Lib.add_anonymous_leaf (inline_extraction (true,[g])); - Lib.add_anonymous_leaf (in_customs (g,[],s))) l + Lib.add_anonymous_leaf (inline_extraction (true,[g])); + Lib.add_anonymous_leaf (in_customs (g,[],s))) l | _ -> error_inductive g diff --git a/plugins/firstorder/formula.ml b/plugins/firstorder/formula.ml index fb363b9393..38dd8992bc 100644 --- a/plugins/firstorder/formula.ml +++ b/plugins/firstorder/formula.ml @@ -41,7 +41,7 @@ let meta_succ m = m+1 let rec nb_prod_after n c= match Constr.kind c with | Prod (_,_,b) ->if n>0 then nb_prod_after (n-1) b else - 1+(nb_prod_after 0 b) + 1+(nb_prod_after 0 b) | _ -> 0 let construct_nhyps env ind = @@ -82,40 +82,40 @@ let kind_of_formula env sigma term = let normalize = special_nf env sigma in let cciterm = special_whd env sigma term in match match_with_imp_term env sigma cciterm with - Some (a,b)-> Arrow (a, pop b) + Some (a,b)-> Arrow (a, pop b) |_-> match match_with_forall_term env sigma cciterm with - Some (_,a,b)-> Forall (a, b) - |_-> + Some (_,a,b)-> Forall (a, b) + |_-> match match_with_nodep_ind env sigma cciterm with - Some (i,l,n)-> - let ind,u=EConstr.destInd sigma i in - let u = EConstr.EInstance.kind sigma u in - let (mib,mip) = Global.lookup_inductive ind in - let nconstr=Array.length mip.mind_consnames in - if Int.equal nconstr 0 then - False((ind,u),l) - else - let has_realargs=(n>0) in - let is_trivial= + Some (i,l,n)-> + let ind,u=EConstr.destInd sigma i in + let u = EConstr.EInstance.kind sigma u in + let (mib,mip) = Global.lookup_inductive ind in + let nconstr=Array.length mip.mind_consnames in + if Int.equal nconstr 0 then + False((ind,u),l) + else + let has_realargs=(n>0) in + let is_trivial= let is_constant n = Int.equal n 0 in Array.exists is_constant mip.mind_consnrealargs in - if Inductiveops.mis_is_recursive (ind,mib,mip) || - (has_realargs && not is_trivial) - then - Atom cciterm - else - if Int.equal nconstr 1 then - And((ind,u),l,is_trivial) - else - Or((ind,u),l,is_trivial) - | _ -> + if Inductiveops.mis_is_recursive (ind,mib,mip) || + (has_realargs && not is_trivial) + then + Atom cciterm + else + if Int.equal nconstr 1 then + And((ind,u),l,is_trivial) + else + Or((ind,u),l,is_trivial) + | _ -> match match_with_sigma_type env sigma cciterm with - Some (i,l)-> + Some (i,l)-> let (ind, u) = EConstr.destInd sigma i in let u = EConstr.EInstance.kind sigma u in Exists((ind, u), l) - |_-> Atom (normalize cciterm) + |_-> Atom (normalize cciterm) type atoms = {positive:constr list;negative:constr list} @@ -132,52 +132,52 @@ let build_atoms env sigma metagen side cciterm = let normalize=special_nf env sigma in let rec build_rec subst polarity cciterm= match kind_of_formula env sigma cciterm with - False(_,_)->if not polarity then trivial:=true + False(_,_)->if not polarity then trivial:=true | Arrow (a,b)-> - build_rec subst (not polarity) a; - build_rec subst polarity b + build_rec subst (not polarity) a; + build_rec subst polarity b | And(i,l,b) | Or(i,l,b)-> - if b then - begin - let unsigned=normalize (substnl subst 0 cciterm) in - if polarity then - positive:= unsigned :: !positive - else - negative:= unsigned :: !negative - end; - let v = ind_hyps env sigma 0 i l in - let g i _ decl = - build_rec subst polarity (lift i (RelDecl.get_type decl)) in - let f l = - List.fold_left_i g (1-(List.length l)) () l in - if polarity && (* we have a constant constructor *) - Array.exists (function []->true|_->false) v - then trivial:=true; - Array.iter f v + if b then + begin + let unsigned=normalize (substnl subst 0 cciterm) in + if polarity then + positive:= unsigned :: !positive + else + negative:= unsigned :: !negative + end; + let v = ind_hyps env sigma 0 i l in + let g i _ decl = + build_rec subst polarity (lift i (RelDecl.get_type decl)) in + let f l = + List.fold_left_i g (1-(List.length l)) () l in + if polarity && (* we have a constant constructor *) + Array.exists (function []->true|_->false) v + then trivial:=true; + Array.iter f v | Exists(i,l)-> - let var=mkMeta (metagen true) in - let v =(ind_hyps env sigma 1 i l).(0) in - let g i _ decl = - build_rec (var::subst) polarity (lift i (RelDecl.get_type decl)) in - List.fold_left_i g (2-(List.length l)) () v + let var=mkMeta (metagen true) in + let v =(ind_hyps env sigma 1 i l).(0) in + let g i _ decl = + build_rec (var::subst) polarity (lift i (RelDecl.get_type decl)) in + List.fold_left_i g (2-(List.length l)) () v | Forall(_,b)-> - let var=mkMeta (metagen true) in - build_rec (var::subst) polarity b + let var=mkMeta (metagen true) in + build_rec (var::subst) polarity b | Atom t-> - let unsigned=substnl subst 0 t in - if not (isMeta sigma unsigned) then (* discarding wildcard atoms *) - if polarity then - positive:= unsigned :: !positive - else - negative:= unsigned :: !negative in + let unsigned=substnl subst 0 t in + if not (isMeta sigma unsigned) then (* discarding wildcard atoms *) + if polarity then + positive:= unsigned :: !positive + else + negative:= unsigned :: !negative in begin match side with - Concl -> build_rec [] true cciterm - | Hyp -> build_rec [] false cciterm - | Hint -> - let rels,head=decompose_prod sigma cciterm in - let subst=List.rev_map (fun _->mkMeta (metagen true)) rels in - build_rec subst false head;trivial:=false (* special for hints *) + Concl -> build_rec [] true cciterm + | Hyp -> build_rec [] false cciterm + | Hint -> + let rels,head=decompose_prod sigma cciterm in + let subst=List.rev_map (fun _->mkMeta (metagen true)) rels in + build_rec subst false head;trivial:=false (* special for hints *) end; (!trivial, {positive= !positive; @@ -209,65 +209,65 @@ type left_pattern= | LA of constr*left_arrow_pattern type t={id:GlobRef.t; - constr:constr; - pat:(left_pattern,right_pattern) sum; - atoms:atoms} + constr:constr; + pat:(left_pattern,right_pattern) sum; + atoms:atoms} let build_formula env sigma side nam typ metagen= let normalize = special_nf env sigma in try let m=meta_succ(metagen false) in let trivial,atoms= - if !qflag then - build_atoms env sigma metagen side typ - else no_atoms in + if !qflag then + build_atoms env sigma metagen side typ + else no_atoms in let pattern= - match side with - Concl -> - let pat= - match kind_of_formula env sigma typ with - False(_,_) -> Rfalse - | Atom a -> raise (Is_atom a) - | And(_,_,_) -> Rand - | Or(_,_,_) -> Ror - | Exists (i,l) -> - let d = RelDecl.get_type (List.last (ind_hyps env sigma 0 i l).(0)) in - Rexists(m,d,trivial) - | Forall (_,a) -> Rforall - | Arrow (a,b) -> Rarrow in - Right pat - | _ -> - let pat= - match kind_of_formula env sigma typ with - False(i,_) -> Lfalse - | Atom a -> raise (Is_atom a) - | And(i,_,b) -> - if b then - let nftyp=normalize typ in raise (Is_atom nftyp) - else Land i - | Or(i,_,b) -> - if b then - let nftyp=normalize typ in raise (Is_atom nftyp) - else Lor i - | Exists (ind,_) -> Lexists ind - | Forall (d,_) -> - Lforall(m,d,trivial) - | Arrow (a,b) -> - let nfa=normalize a in - LA (nfa, - match kind_of_formula env sigma a with - False(i,l)-> LLfalse(i,l) - | Atom t-> LLatom - | And(i,l,_)-> LLand(i,l) - | Or(i,l,_)-> LLor(i,l) - | Arrow(a,c)-> LLarrow(a,c,b) - | Exists(i,l)->LLexists(i,l) - | Forall(_,_)->LLforall a) in - Left pat + match side with + Concl -> + let pat= + match kind_of_formula env sigma typ with + False(_,_) -> Rfalse + | Atom a -> raise (Is_atom a) + | And(_,_,_) -> Rand + | Or(_,_,_) -> Ror + | Exists (i,l) -> + let d = RelDecl.get_type (List.last (ind_hyps env sigma 0 i l).(0)) in + Rexists(m,d,trivial) + | Forall (_,a) -> Rforall + | Arrow (a,b) -> Rarrow in + Right pat + | _ -> + let pat= + match kind_of_formula env sigma typ with + False(i,_) -> Lfalse + | Atom a -> raise (Is_atom a) + | And(i,_,b) -> + if b then + let nftyp=normalize typ in raise (Is_atom nftyp) + else Land i + | Or(i,_,b) -> + if b then + let nftyp=normalize typ in raise (Is_atom nftyp) + else Lor i + | Exists (ind,_) -> Lexists ind + | Forall (d,_) -> + Lforall(m,d,trivial) + | Arrow (a,b) -> + let nfa=normalize a in + LA (nfa, + match kind_of_formula env sigma a with + False(i,l)-> LLfalse(i,l) + | Atom t-> LLatom + | And(i,l,_)-> LLand(i,l) + | Or(i,l,_)-> LLor(i,l) + | Arrow(a,c)-> LLarrow(a,c,b) + | Exists(i,l)->LLexists(i,l) + | Forall(_,_)->LLforall a) in + Left pat in - Left {id=nam; - constr=normalize typ; - pat=pattern; - atoms=atoms} + Left {id=nam; + constr=normalize typ; + pat=pattern; + atoms=atoms} with Is_atom a-> Right a (* already in nf *) diff --git a/plugins/firstorder/formula.mli b/plugins/firstorder/formula.mli index dc422fa284..b8a619d1e6 100644 --- a/plugins/firstorder/formula.mli +++ b/plugins/firstorder/formula.mli @@ -66,9 +66,9 @@ type left_pattern= | LA of constr*left_arrow_pattern type t={id: GlobRef.t; - constr: constr; - pat: (left_pattern,right_pattern) sum; - atoms: atoms} + constr: constr; + pat: (left_pattern,right_pattern) sum; + atoms: atoms} (*exception Is_atom of constr*) diff --git a/plugins/firstorder/g_ground.mlg b/plugins/firstorder/g_ground.mlg index 35cd10a1ff..2bc79d45d4 100644 --- a/plugins/firstorder/g_ground.mlg +++ b/plugins/firstorder/g_ground.mlg @@ -41,7 +41,7 @@ let ()= optread=(fun ()->Some !ground_depth); optwrite= (function - None->ground_depth:=3 + None->ground_depth:=3 | Some i->ground_depth:=(max i 0))} in declare_int_option gdopt @@ -68,7 +68,7 @@ let default_intuition_tac = Tacenv.register_ml_tactic name [| tac |]; Tacexpr.TacML (CAst.make (entry, [])) -let (set_default_solver, default_solver, print_default_solver) = +let (set_default_solver, default_solver, print_default_solver) = Tactic_option.declare_tactic_option ~default:default_intuition_tac "Firstorder default solver" } @@ -95,12 +95,12 @@ let gen_ground_tac flag taco ids bases = Proofview.Goal.enter begin fun gl -> qflag:=flag; let solver= - match taco with - Some tac-> tac - | None-> snd (default_solver ()) in + match taco with + Some tac-> tac + | None-> snd (default_solver ()) in let startseq k = Proofview.Goal.enter begin fun gl -> - let seq=empty_seq !ground_depth in + let seq=empty_seq !ground_depth in let seq, sigma = extend_with_ref_list (pf_env gl) (project gl) ids seq in let seq, sigma = extend_with_auto_hints (pf_env gl) (project gl) bases seq in tclTHEN (Proofview.Unsafe.tclEVARS sigma) (k seq) @@ -124,10 +124,10 @@ let defined_connectives=lazy let normalize_evaluables= onAllHypsAndConcl (function - None->unfold_in_concl (Lazy.force defined_connectives) + None->unfold_in_concl (Lazy.force defined_connectives) | Some id-> - unfold_in_hyp (Lazy.force defined_connectives) - (Tacexpr.InHypType id)) *) + unfold_in_hyp (Lazy.force defined_connectives) + (Tacexpr.InHypType id)) *) open Ppconstr open Printer diff --git a/plugins/firstorder/ground.ml b/plugins/firstorder/ground.ml index e134562702..2f26226f4e 100644 --- a/plugins/firstorder/ground.ml +++ b/plugins/firstorder/ground.ml @@ -41,89 +41,89 @@ let ground_tac solver startseq = in tclORELSE (axiom_tac seq.gl seq) begin - try - let (hd,seq1)=take_formula (project gl) seq - and re_add s=re_add_formula_list (project gl) skipped s in - let continue=toptac [] - and backtrack =toptac (hd::skipped) seq1 in - match hd.pat with - Right rpat-> - begin - match rpat with - Rand-> - and_tac backtrack continue (re_add seq1) - | Rforall-> - let backtrack1= - if !qflag then - tclFAIL 0 (Pp.str "reversible in 1st order mode") - else - backtrack in - forall_tac backtrack1 continue (re_add seq1) - | Rarrow-> - arrow_tac backtrack continue (re_add seq1) - | Ror-> - or_tac backtrack continue (re_add seq1) - | Rfalse->backtrack - | Rexists(i,dom,triv)-> - let (lfp,seq2)=collect_quantified (project gl) seq in - let backtrack2=toptac (lfp@skipped) seq2 in - if !qflag && seq.depth>0 then - quantified_tac lfp backtrack2 - continue (re_add seq) - else - backtrack2 (* need special backtracking *) - end - | Left lpat-> - begin - match lpat with - Lfalse-> - left_false_tac hd.id - | Land ind-> - left_and_tac ind backtrack - hd.id continue (re_add seq1) - | Lor ind-> - left_or_tac ind backtrack - hd.id continue (re_add seq1) - | Lforall (_,_,_)-> - let (lfp,seq2)=collect_quantified (project gl) seq in - let backtrack2=toptac (lfp@skipped) seq2 in - if !qflag && seq.depth>0 then - quantified_tac lfp backtrack2 - continue (re_add seq) - else - backtrack2 (* need special backtracking *) - | Lexists ind -> - if !qflag then - left_exists_tac ind backtrack hd.id - continue (re_add seq1) - else backtrack - | LA (typ,lap)-> - let la_tac= - begin - match lap with - LLatom -> backtrack - | LLand (ind,largs) | LLor(ind,largs) - | LLfalse (ind,largs)-> - (ll_ind_tac ind largs backtrack - hd.id continue (re_add seq1)) - | LLforall p -> - if seq.depth>0 && !qflag then - (ll_forall_tac p backtrack - hd.id continue (re_add seq1)) - else backtrack - | LLexists (ind,l) -> - if !qflag then - ll_ind_tac ind l backtrack - hd.id continue (re_add seq1) - else - backtrack - | LLarrow (a,b,c) -> - (ll_arrow_tac a b c backtrack - hd.id continue (re_add seq1)) - end in - ll_atom_tac typ la_tac hd.id continue (re_add seq1) - end - with Heap.EmptyHeap->solver + try + let (hd,seq1)=take_formula (project gl) seq + and re_add s=re_add_formula_list (project gl) skipped s in + let continue=toptac [] + and backtrack =toptac (hd::skipped) seq1 in + match hd.pat with + Right rpat-> + begin + match rpat with + Rand-> + and_tac backtrack continue (re_add seq1) + | Rforall-> + let backtrack1= + if !qflag then + tclFAIL 0 (Pp.str "reversible in 1st order mode") + else + backtrack in + forall_tac backtrack1 continue (re_add seq1) + | Rarrow-> + arrow_tac backtrack continue (re_add seq1) + | Ror-> + or_tac backtrack continue (re_add seq1) + | Rfalse->backtrack + | Rexists(i,dom,triv)-> + let (lfp,seq2)=collect_quantified (project gl) seq in + let backtrack2=toptac (lfp@skipped) seq2 in + if !qflag && seq.depth>0 then + quantified_tac lfp backtrack2 + continue (re_add seq) + else + backtrack2 (* need special backtracking *) + end + | Left lpat-> + begin + match lpat with + Lfalse-> + left_false_tac hd.id + | Land ind-> + left_and_tac ind backtrack + hd.id continue (re_add seq1) + | Lor ind-> + left_or_tac ind backtrack + hd.id continue (re_add seq1) + | Lforall (_,_,_)-> + let (lfp,seq2)=collect_quantified (project gl) seq in + let backtrack2=toptac (lfp@skipped) seq2 in + if !qflag && seq.depth>0 then + quantified_tac lfp backtrack2 + continue (re_add seq) + else + backtrack2 (* need special backtracking *) + | Lexists ind -> + if !qflag then + left_exists_tac ind backtrack hd.id + continue (re_add seq1) + else backtrack + | LA (typ,lap)-> + let la_tac= + begin + match lap with + LLatom -> backtrack + | LLand (ind,largs) | LLor(ind,largs) + | LLfalse (ind,largs)-> + (ll_ind_tac ind largs backtrack + hd.id continue (re_add seq1)) + | LLforall p -> + if seq.depth>0 && !qflag then + (ll_forall_tac p backtrack + hd.id continue (re_add seq1)) + else backtrack + | LLexists (ind,l) -> + if !qflag then + ll_ind_tac ind l backtrack + hd.id continue (re_add seq1) + else + backtrack + | LLarrow (a,b,c) -> + (ll_arrow_tac a b c backtrack + hd.id continue (re_add seq1)) + end in + ll_atom_tac typ la_tac hd.id continue (re_add seq1) + end + with Heap.EmptyHeap->solver end end in let n = List.length (Proofview.Goal.hyps gl) in diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml index eff0db5bf4..e131cad7da 100644 --- a/plugins/firstorder/instances.ml +++ b/plugins/firstorder/instances.ml @@ -26,13 +26,13 @@ open Context.Rel.Declaration let compare_instance inst1 inst2= let cmp c1 c2 = Constr.compare (EConstr.Unsafe.to_constr c1) (EConstr.Unsafe.to_constr c2) in - match inst1,inst2 with - Phantom(d1),Phantom(d2)-> - (cmp d1 d2) - | Real((m1,c1),n1),Real((m2,c2),n2)-> - ((-) =? (-) ==? cmp) m2 m1 n1 n2 c1 c2 - | Phantom(_),Real((m,_),_)-> if Int.equal m 0 then -1 else 1 - | Real((m,_),_),Phantom(_)-> if Int.equal m 0 then 1 else -1 + match inst1,inst2 with + Phantom(d1),Phantom(d2)-> + (cmp d1 d2) + | Real((m1,c1),n1),Real((m2,c2),n2)-> + ((-) =? (-) ==? cmp) m2 m1 n1 n2 c1 c2 + | Phantom(_),Real((m,_),_)-> if Int.equal m 0 then -1 else 1 + | Real((m,_),_),Phantom(_)-> if Int.equal m 0 then 1 else -1 let compare_gr id1 id2 = if id1==id2 then 0 else @@ -53,7 +53,7 @@ module IS=Set.Make(OrderedInstance) let make_simple_atoms seq= let ratoms= match seq.glatom with - Some t->[t] + Some t->[t] | None->[] in {negative=seq.latoms;positive=ratoms} @@ -63,9 +63,9 @@ let do_sequent sigma setref triv id seq i dom atoms= let do_atoms a1 a2 = let do_pair t1 t2 = match unif_atoms sigma i dom t1 t2 with - None->() - | Some (Phantom _) ->phref:=true - | Some c ->flag:=false;setref:=IS.add (c,id) !setref in + None->() + | Some (Phantom _) ->phref:=true + | Some c ->flag:=false;setref:=IS.add (c,id) !setref in List.iter (fun t->List.iter (do_pair t) a2.negative) a1.positive; List.iter (fun t->List.iter (do_pair t) a2.positive) a1.negative in HP.iter (fun lf->do_atoms atoms lf.atoms) seq.redexes; @@ -75,8 +75,8 @@ let do_sequent sigma setref triv id seq i dom atoms= let match_one_quantified_hyp sigma setref seq lf= match lf.pat with Left(Lforall(i,dom,triv))|Right(Rexists(i,dom,triv))-> - if do_sequent sigma setref triv lf.id seq i dom lf.atoms then - setref:=IS.add ((Phantom dom),lf.id) !setref + if do_sequent sigma setref triv lf.id seq i dom lf.atoms then + setref:=IS.add ((Phantom dom),lf.id) !setref | _ -> anomaly (Pp.str "can't happen.") let give_instances sigma lf seq= @@ -90,10 +90,10 @@ let rec collect_quantified sigma seq= try let hd,seq1=take_formula sigma seq in (match hd.pat with - Left(Lforall(_,_,_)) | Right(Rexists(_,_,_)) -> - let (q,seq2)=collect_quantified sigma seq1 in - ((hd::q),seq2) - | _->[],seq) + Left(Lforall(_,_,_)) | Right(Rexists(_,_,_)) -> + let (q,seq2)=collect_quantified sigma seq1 in + ((hd::q),seq2) + | _->[],seq) with Heap.EmptyHeap -> [],seq (* open instances processor *) @@ -104,19 +104,19 @@ let mk_open_instance env evmap id idc m t = let var_id= if id==dummy_id then dummy_bvid else let typ=Typing.unsafe_type_of env evmap idc in - (* since we know we will get a product, - reduction is not too expensive *) + (* since we know we will get a product, + reduction is not too expensive *) let (nam,_,_)=destProd evmap (whd_all env evmap typ) in match nam.Context.binder_name with - Name id -> id - | Anonymous -> dummy_bvid in + Name id -> id + | Anonymous -> dummy_bvid in let revt=substl (List.init m (fun i->mkRel (m-i))) t in let rec aux n avoid env evmap decls = if Int.equal n 0 then evmap, decls else let nid=(fresh_id_in_env avoid var_id env) in let (evmap, (c, _)) = Evarutil.new_type_evar env evmap Evd.univ_flexible in let decl = LocalAssum (Context.make_annot (Name nid) Sorts.Relevant, c) in - aux (n-1) (Id.Set.add nid avoid) (EConstr.push_rel decl env) evmap (decl::decls) in + aux (n-1) (Id.Set.add nid avoid) (EConstr.push_rel decl env) evmap (decl::decls) in let evmap, decls = aux m Id.Set.empty env evmap [] in (evmap, decls, revt) @@ -128,49 +128,49 @@ let left_instance_tac (inst,id) continue seq= let sigma = project gl in match inst with Phantom dom-> - if lookup sigma (id,None) seq then - tclFAIL 0 (Pp.str "already done") - else - tclTHENS (cut dom) - [tclTHENLIST - [introf; + if lookup sigma (id,None) seq then + tclFAIL 0 (Pp.str "already done") + else + tclTHENS (cut dom) + [tclTHENLIST + [introf; (pf_constr_of_global id >>= fun idc -> Proofview.Goal.enter begin fun gl -> let id0 = List.nth (pf_ids_of_hyps gl) 0 in generalize [mkApp(idc, [|mkVar id0|])] end); - introf; - tclSOLVE [wrap 1 false continue - (deepen (record (id,None) seq))]]; - tclTRY assumption] + introf; + tclSOLVE [wrap 1 false continue + (deepen (record (id,None) seq))]]; + tclTRY assumption] | Real((m,t),_)-> let c = (m, EConstr.to_constr sigma t) in - if lookup sigma (id,Some c) seq then - tclFAIL 0 (Pp.str "already done") - else - let special_generalize= - if m>0 then - (pf_constr_of_global id >>= fun idc -> - Proofview.Goal.enter begin fun gl-> - let (evmap, rc, ot) = mk_open_instance (pf_env gl) (project gl) id idc m t in - let gt= - it_mkLambda_or_LetIn - (mkApp(idc,[|ot|])) rc in - let evmap, _ = - try Typing.type_of (pf_env gl) evmap gt - with e when CErrors.noncritical e -> - user_err Pp.(str "Untypable instance, maybe higher-order non-prenex quantification") in + if lookup sigma (id,Some c) seq then + tclFAIL 0 (Pp.str "already done") + else + let special_generalize= + if m>0 then + (pf_constr_of_global id >>= fun idc -> + Proofview.Goal.enter begin fun gl-> + let (evmap, rc, ot) = mk_open_instance (pf_env gl) (project gl) id idc m t in + let gt= + it_mkLambda_or_LetIn + (mkApp(idc,[|ot|])) rc in + let evmap, _ = + try Typing.type_of (pf_env gl) evmap gt + with e when CErrors.noncritical e -> + user_err Pp.(str "Untypable instance, maybe higher-order non-prenex quantification") in Proofview.tclTHEN (Proofview.Unsafe.tclEVARS evmap) - (generalize [gt]) + (generalize [gt]) end) - else - pf_constr_of_global id >>= fun idc -> generalize [mkApp(idc,[|t|])] - in - tclTHENLIST - [special_generalize; - introf; - tclSOLVE - [wrap 1 false continue (deepen (record (id,Some c) seq))]] + else + pf_constr_of_global id >>= fun idc -> generalize [mkApp(idc,[|t|])] + in + tclTHENLIST + [special_generalize; + introf; + tclSOLVE + [wrap 1 false continue (deepen (record (id,Some c) seq))]] end let right_instance_tac inst continue seq= @@ -178,20 +178,20 @@ let right_instance_tac inst continue seq= Proofview.Goal.enter begin fun gl -> match inst with Phantom dom -> - tclTHENS (cut dom) - [tclTHENLIST - [introf; + tclTHENS (cut dom) + [tclTHENLIST + [introf; Proofview.Goal.enter begin fun gl -> let id0 = List.nth (pf_ids_of_hyps gl) 0 in split (Tactypes.ImplicitBindings [mkVar id0]) end; - tclSOLVE [wrap 0 true continue (deepen seq)]]; - tclTRY assumption] + tclSOLVE [wrap 0 true continue (deepen seq)]]; + tclTRY assumption] | Real ((0,t),_) -> (tclTHEN (split (Tactypes.ImplicitBindings [t])) - (tclSOLVE [wrap 0 true continue (deepen seq)])) + (tclSOLVE [wrap 0 true continue (deepen seq)])) | Real ((m,t),_) -> - tclFAIL 0 (Pp.str "not implemented ... yet") + tclFAIL 0 (Pp.str "not implemented ... yet") end let instance_tac inst= diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml index 79386f7ac9..3413db930b 100644 --- a/plugins/firstorder/rules.ml +++ b/plugins/firstorder/rules.ml @@ -40,13 +40,13 @@ let wrap n b continue seq = let rec aux i nc ctx= if i<=0 then seq else match nc with - []->anomaly (Pp.str "Not the expected number of hyps.") - | nd::q-> + []->anomaly (Pp.str "Not the expected number of hyps.") + | nd::q-> let id = NamedDecl.get_id nd in - if occur_var env sigma id (pf_concl gls) || - List.exists (occur_var_in_decl env sigma id) ctx then - (aux (i-1) q (nd::ctx)) - else + if occur_var env sigma id (pf_concl gls) || + List.exists (occur_var_in_decl env sigma id) ctx then + (aux (i-1) q (nd::ctx)) + else add_formula env sigma Hyp (GlobRef.VarRef id) (NamedDecl.get_type nd) (aux (i-1) q (nd::ctx)) in let seq1=aux n nc [] in let seq2=if b then @@ -72,17 +72,17 @@ let ll_atom_tac a backtrack id continue seq = let open EConstr in tclIFTHENELSE (tclTHENLIST - [(Proofview.tclEVARMAP >>= fun sigma -> + [(Proofview.tclEVARMAP >>= fun sigma -> let gr = try Proofview.tclUNIT (find_left sigma a seq) with Not_found -> tclFAIL 0 (Pp.str "No link") in gr >>= fun gr -> pf_constr_of_global gr >>= fun left -> - pf_constr_of_global id >>= fun id -> - generalize [(mkApp(id, [|left|]))]); - clear_global id; - intro]) + pf_constr_of_global id >>= fun id -> + generalize [(mkApp(id, [|left|]))]); + clear_global id; + intro]) (wrap 1 false continue seq) backtrack (* right connectives rules *) @@ -151,12 +151,12 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq = EConstr.it_mkLambda_or_LetIn head rc in let lp=Array.length rcs in let newhyps idc =List.init lp (myterm idc) in - tclIFTHENELSE - (tclTHENLIST - [(pf_constr_of_global id >>= fun idc -> generalize (newhyps idc)); - clear_global id; - tclDO lp intro]) - (wrap lp false continue seq) backtrack + tclIFTHENELSE + (tclTHENLIST + [(pf_constr_of_global id >>= fun idc -> generalize (newhyps idc)); + clear_global id; + tclDO lp intro]) + (wrap lp false continue seq) backtrack end let ll_arrow_tac a b c backtrack id continue seq= @@ -167,18 +167,18 @@ let ll_arrow_tac a b c backtrack id continue seq= mkApp (idc, [|mkLambda (Context.make_annot Anonymous Sorts.Relevant,(lift 1 a),(mkRel 2))|])) in tclORELSE (tclTHENS (cut c) - [tclTHENLIST - [introf; - clear_global id; - wrap 1 false continue seq]; - tclTHENS (cut cc) + [tclTHENLIST + [introf; + clear_global id; + wrap 1 false continue seq]; + tclTHENS (cut cc) [(pf_constr_of_global id >>= fun c -> exact_no_check c); - tclTHENLIST - [(pf_constr_of_global id >>= fun idc -> generalize [d idc]); - clear_global id; - introf; - introf; - tclCOMPLETE (wrap 2 true continue seq)]]]) + tclTHENLIST + [(pf_constr_of_global id >>= fun idc -> generalize [d idc]); + clear_global id; + introf; + introf; + tclCOMPLETE (wrap 2 true continue seq)]]]) backtrack (* quantifier rules (easy side) *) @@ -187,8 +187,8 @@ let forall_tac backtrack continue seq= tclORELSE (tclIFTHENELSE intro (wrap 0 true continue seq) (tclORELSE - (tclTHEN introf (tclCOMPLETE (wrap 0 true continue seq))) - backtrack)) + (tclTHEN introf (tclCOMPLETE (wrap 0 true continue seq))) + backtrack)) (if !qflag then tclFAIL 0 (Pp.str "reversible in 1st order mode") else @@ -209,18 +209,18 @@ let ll_forall_tac prod backtrack id continue seq= tclORELSE (tclTHENS (cut prod) [tclTHENLIST - [intro; + [intro; (pf_constr_of_global id >>= fun idc -> - Proofview.Goal.enter begin fun gls-> + Proofview.Goal.enter begin fun gls-> let open EConstr in - let id0 = List.nth (pf_ids_of_hyps gls) 0 in + let id0 = List.nth (pf_ids_of_hyps gls) 0 in let term=mkApp(idc,[|mkVar(id0)|]) in tclTHEN (generalize [term]) (clear [id0]) end); - clear_global id; - intro; - tclCOMPLETE (wrap 1 false continue (deepen seq))]; - tclCOMPLETE (wrap 0 true continue (deepen seq))]) + clear_global id; + intro; + tclCOMPLETE (wrap 1 false continue (deepen seq))]; + tclCOMPLETE (wrap 0 true continue (deepen seq))]) backtrack (* rules for instantiation with unification moved to instances.ml *) diff --git a/plugins/firstorder/sequent.ml b/plugins/firstorder/sequent.ml index e53412383c..9ff05c33e4 100644 --- a/plugins/firstorder/sequent.ml +++ b/plugins/firstorder/sequent.ml @@ -23,37 +23,37 @@ let newcnt ()= let priority = (* pure heuristics, <=0 for non reversible *) function Right rf-> - begin - match rf with - Rarrow -> 100 - | Rand -> 40 - | Ror -> -15 - | Rfalse -> -50 - | Rforall -> 100 - | Rexists (_,_,_) -> -29 - end + begin + match rf with + Rarrow -> 100 + | Rand -> 40 + | Ror -> -15 + | Rfalse -> -50 + | Rforall -> 100 + | Rexists (_,_,_) -> -29 + end | Left lf -> - match lf with - Lfalse -> 999 - | Land _ -> 90 - | Lor _ -> 40 - | Lforall (_,_,_) -> -30 - | Lexists _ -> 60 - | LA(_,lap) -> - match lap with - LLatom -> 0 - | LLfalse (_,_) -> 100 - | LLand (_,_) -> 80 - | LLor (_,_) -> 70 - | LLforall _ -> -20 - | LLexists (_,_) -> 50 - | LLarrow (_,_,_) -> -10 + match lf with + Lfalse -> 999 + | Land _ -> 90 + | Lor _ -> 40 + | Lforall (_,_,_) -> -30 + | Lexists _ -> 60 + | LA(_,lap) -> + match lap with + LLatom -> 0 + | LLfalse (_,_) -> 100 + | LLand (_,_) -> 80 + | LLor (_,_) -> 70 + | LLforall _ -> -20 + | LLexists (_,_) -> 50 + | LLarrow (_,_,_) -> -10 module OrderedFormula= struct type t=Formula.t let compare e1 e2= - (priority e1.pat) - (priority e2.pat) + (priority e1.pat) - (priority e2.pat) end type h_item = GlobRef.t * (int*Constr.t) option @@ -89,8 +89,8 @@ let cm_remove sigma typ nam cm= let l=CM.find typ cm in let l0=List.filter (fun id-> not (GlobRef.equal id nam)) l in match l0 with - []->CM.remove typ cm - | _ ->CM.add typ l0 cm + []->CM.remove typ cm + | _ ->CM.add typ l0 cm with Not_found ->cm module HP=Heap.Functional(OrderedFormula) @@ -114,35 +114,35 @@ let lookup sigma item seq= match item with (_,None)->false | (id,Some (m, t))-> - let p (id2,o)= - match o with - None -> false + let p (id2,o)= + match o with + None -> false | Some (m2, t2)-> GlobRef.equal id id2 && m2>m && more_general sigma (m2, EConstr.of_constr t2) (m, EConstr.of_constr t) in - History.exists p seq.history + History.exists p seq.history let add_formula env sigma side nam t seq = match build_formula env sigma side nam t seq.cnt with Left f-> - begin - match side with - Concl -> - {seq with - redexes=HP.add f seq.redexes; - gl=f.constr; - glatom=None} - | _ -> - {seq with - redexes=HP.add f seq.redexes; - context=cm_add sigma f.constr nam seq.context} - end + begin + match side with + Concl -> + {seq with + redexes=HP.add f seq.redexes; + gl=f.constr; + glatom=None} + | _ -> + {seq with + redexes=HP.add f seq.redexes; + context=cm_add sigma f.constr nam seq.context} + end | Right t-> - match side with - Concl -> - {seq with gl=t;glatom=Some t} - | _ -> - {seq with - context=cm_add sigma t nam seq.context; - latoms=t::seq.latoms} + match side with + Concl -> + {seq with gl=t;glatom=Some t} + | _ -> + {seq with + context=cm_add sigma t nam seq.context; + latoms=t::seq.latoms} let re_add_formula_list sigma lf seq= let do_one f cm= @@ -166,14 +166,14 @@ let rec take_formula sigma seq= and hp=HP.remove seq.redexes in if hd.id == dummy_id then let nseq={seq with redexes=hp} in - if seq.gl==hd.constr then - hd,nseq - else - take_formula sigma nseq (* discarding deprecated goal *) + if seq.gl==hd.constr then + hd,nseq + else + take_formula sigma nseq (* discarding deprecated goal *) else hd,{seq with - redexes=hp; - context=cm_remove sigma hd.constr hd.id seq.context} + redexes=hp; + context=cm_remove sigma hd.constr hd.id seq.context} let empty_seq depth= {redexes=HP.empty; @@ -191,7 +191,7 @@ let expand_constructor_hints = List.init (Inductiveops.nconstructors (Global.env()) ind) (fun i -> GlobRef.ConstructRef (ind,i+1)) | gr -> - [gr]) + [gr]) let extend_with_ref_list env sigma l seq = let l = expand_constructor_hints l in @@ -207,22 +207,22 @@ let extend_with_auto_hints env sigma l seq = let seqref=ref seq in let f p_a_t = match repr_hint p_a_t.code with - Res_pf (c,_) | Give_exact (c,_) + Res_pf (c,_) | Give_exact (c,_) | Res_pf_THEN_trivial_fail (c,_) -> let (c, _, _) = c in - (try - let (gr, _) = Termops.global_of_constr sigma c in - let typ=(Typing.unsafe_type_of env sigma c) in - seqref:=add_formula env sigma Hint gr typ !seqref - with Not_found->()) + (try + let (gr, _) = Termops.global_of_constr sigma c in + let typ=(Typing.unsafe_type_of env sigma c) in + seqref:=add_formula env sigma Hint gr typ !seqref + with Not_found->()) | _-> () in let g _ _ l = List.iter f l in let h dbname= let hdb= try - searchtable_map dbname + searchtable_map dbname with Not_found-> - user_err Pp.(str ("Firstorder: "^dbname^" : No such Hint database")) in + user_err Pp.(str ("Firstorder: "^dbname^" : No such Hint database")) in Hint_db.iter g hdb in List.iter h l; !seqref, sigma (*FIXME: forgetting about universes*) @@ -239,9 +239,9 @@ let print_cmap map= cut () ++ s in (v 0 - (str "-----" ++ - cut () ++ - CM.fold print_entry map (mt ()) ++ - str "-----")) + (str "-----" ++ + cut () ++ + CM.fold print_entry map (mt ()) ++ + str "-----")) diff --git a/plugins/firstorder/sequent.mli b/plugins/firstorder/sequent.mli index 724e1abcc4..2e262fd996 100644 --- a/plugins/firstorder/sequent.mli +++ b/plugins/firstorder/sequent.mli @@ -28,12 +28,12 @@ module HP: Heap.S with type elt=Formula.t type t = {redexes:HP.t; context: GlobRef.t list CM.t; - latoms:constr list; - gl:types; - glatom:constr option; - cnt:counter; - history:History.t; - depth:int} + latoms:constr list; + gl:types; + glatom:constr option; + cnt:counter; + history:History.t; + depth:int} val deepen: t -> t diff --git a/plugins/firstorder/unify.ml b/plugins/firstorder/unify.ml index 35b64ccb8f..6fa831fda9 100644 --- a/plugins/firstorder/unify.ml +++ b/plugins/firstorder/unify.ml @@ -38,58 +38,58 @@ let unif evd t1 t2= let rec head_reduce t= (* forbids non-sigma-normal meta in head position*) match EConstr.kind evd t with - Meta i-> - (try - head_reduce (Int.List.assoc i !sigma) - with Not_found->t) + Meta i-> + (try + head_reduce (Int.List.assoc i !sigma) + with Not_found->t) | _->t in Queue.add (t1,t2) bige; try while true do let t1,t2=Queue.take bige in let nt1=head_reduce (whd_betaiotazeta evd t1) and nt2=head_reduce (whd_betaiotazeta evd t2) in - match (EConstr.kind evd nt1),(EConstr.kind evd nt2) with - Meta i,Meta j-> - if not (Int.equal i j) then - if i<j then bind j nt1 - else bind i nt2 - | Meta i,_ -> - let t=subst_meta !sigma nt2 in - if Int.Set.is_empty (free_rels evd t) && + match (EConstr.kind evd nt1),(EConstr.kind evd nt2) with + Meta i,Meta j-> + if not (Int.equal i j) then + if i<j then bind j nt1 + else bind i nt2 + | Meta i,_ -> + let t=subst_meta !sigma nt2 in + if Int.Set.is_empty (free_rels evd t) && not (occur_metavariable evd i t) then - bind i t else raise (UFAIL(nt1,nt2)) - | _,Meta i -> - let t=subst_meta !sigma nt1 in - if Int.Set.is_empty (free_rels evd t) && + bind i t else raise (UFAIL(nt1,nt2)) + | _,Meta i -> + let t=subst_meta !sigma nt1 in + if Int.Set.is_empty (free_rels evd t) && not (occur_metavariable evd i t) then - bind i t else raise (UFAIL(nt1,nt2)) - | Cast(_,_,_),_->Queue.add (strip_outer_cast evd nt1,nt2) bige - | _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast evd nt2) bige + bind i t else raise (UFAIL(nt1,nt2)) + | Cast(_,_,_),_->Queue.add (strip_outer_cast evd nt1,nt2) bige + | _,Cast(_,_,_)->Queue.add (nt1,strip_outer_cast evd nt2) bige | (Prod(_,a,b),Prod(_,c,d))|(Lambda(_,a,b),Lambda(_,c,d))-> - Queue.add (a,c) bige;Queue.add (pop b,pop d) bige - | Case (_,pa,ca,va),Case (_,pb,cb,vb)-> - Queue.add (pa,pb) bige; - Queue.add (ca,cb) bige; - let l=Array.length va in + Queue.add (a,c) bige;Queue.add (pop b,pop d) bige + | Case (_,pa,ca,va),Case (_,pb,cb,vb)-> + Queue.add (pa,pb) bige; + Queue.add (ca,cb) bige; + let l=Array.length va in if not (Int.equal l (Array.length vb)) then - raise (UFAIL (nt1,nt2)) - else - for i=0 to l-1 do - Queue.add (va.(i),vb.(i)) bige - done - | App(ha,va),App(hb,vb)-> - Queue.add (ha,hb) bige; - let l=Array.length va in - if not (Int.equal l (Array.length vb)) then - raise (UFAIL (nt1,nt2)) - else - for i=0 to l-1 do - Queue.add (va.(i),vb.(i)) bige - done - | _->if not (eq_constr_nounivs evd nt1 nt2) then raise (UFAIL (nt1,nt2)) + raise (UFAIL (nt1,nt2)) + else + for i=0 to l-1 do + Queue.add (va.(i),vb.(i)) bige + done + | App(ha,va),App(hb,vb)-> + Queue.add (ha,hb) bige; + let l=Array.length va in + if not (Int.equal l (Array.length vb)) then + raise (UFAIL (nt1,nt2)) + else + for i=0 to l-1 do + Queue.add (va.(i),vb.(i)) bige + done + | _->if not (eq_constr_nounivs evd nt1 nt2) then raise (UFAIL (nt1,nt2)) done; assert false - (* this place is unreachable but needed for the sake of typing *) + (* this place is unreachable but needed for the sake of typing *) with Queue.Empty-> !sigma let value evd i t= @@ -99,7 +99,7 @@ let value evd i t= if isMeta evd term && Int.equal (destMeta evd term) i then 0 else let f v t=add v (vaux t) in let vr=EConstr.fold evd f (-1) term in - if vr<0 then -1 else vr+1 in + if vr<0 then -1 else vr+1 in vaux t type instance= @@ -111,14 +111,14 @@ let mk_rel_inst evd t= let rel_env=ref [] in let rec renum_rec d t= match EConstr.kind evd t with - Meta n-> - (try - mkRel (d+(Int.List.assoc n !rel_env)) - with Not_found-> - let m= !new_rel in - incr new_rel; - rel_env:=(n,m) :: !rel_env; - mkRel (m+d)) + Meta n-> + (try + mkRel (d+(Int.List.assoc n !rel_env)) + with Not_found-> + let m= !new_rel in + incr new_rel; + rel_env:=(n,m) :: !rel_env; + mkRel (m+d)) | _ -> EConstr.map_with_binders evd succ renum_rec d t in let nt=renum_rec 0 t in (!new_rel - 1,nt) @@ -142,5 +142,5 @@ let more_general evd (m1,t1) (m2,t2)= try let sigma=unif evd mt1 mt2 in let p (n,t)= n<m1 || isMeta evd t in - List.for_all p sigma + List.for_all p sigma with UFAIL(_,_)->false diff --git a/plugins/funind/g_indfun.mlg b/plugins/funind/g_indfun.mlg index 2b990400e3..a02cb24bee 100644 --- a/plugins/funind/g_indfun.mlg +++ b/plugins/funind/g_indfun.mlg @@ -99,9 +99,9 @@ TACTIC EXTEND newfunind | ["functional" "induction" ne_constr_list(cl) fun_ind_using(princl) with_names(pat)] -> { let c = match cl with - | [] -> assert false - | [c] -> c - | c::cl -> EConstr.applist(c,cl) + | [] -> assert false + | [c] -> c + | c::cl -> EConstr.applist(c,cl) in Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl } END @@ -110,9 +110,9 @@ TACTIC EXTEND snewfunind | ["soft" "functional" "induction" ne_constr_list(cl) fun_ind_using(princl) with_names(pat)] -> { let c = match cl with - | [] -> assert false - | [c] -> c - | c::cl -> EConstr.applist(c,cl) + | [] -> assert false + | [c] -> c + | c::cl -> EConstr.applist(c,cl) in Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl } END @@ -260,7 +260,7 @@ VERNAC COMMAND EXTEND NewFunctionalScheme let names = List.map (fun (_,na,_) -> na) fas in warning_error names e end - | _ -> assert false (* we can only have non empty list *) + | _ -> assert false (* we can only have non empty list *) end | e when CErrors.noncritical e -> let names = List.map (fun (_,na,_) -> na) fas in diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml index 895b6a37ee..e41b92d4dc 100644 --- a/plugins/funind/glob_term_to_relation.ml +++ b/plugins/funind/glob_term_to_relation.ml @@ -38,7 +38,7 @@ let rec solve_trivial_holes pat_as_term e = | GApp(fp,argsp),GApp(fe,argse) when glob_constr_eq fp fe -> DAst.make (GApp((solve_trivial_holes fp fe),List.map2 solve_trivial_holes argsp argse)) | _,_ -> pat_as_term - + (* compose_glob_context [(bt_1,n_1,t_1);......] rt returns b_1(n_1,t_1,.....,bn(n_k,t_k,rt)) where the b_i's are the @@ -90,11 +90,11 @@ let combine_results = let pre_result = List.map ( fun res1 -> (* for each result in arg_res *) - List.map (* we add it in each args_res *) - (fun res2 -> - combine_fun res1 res2 - ) - res2.result + List.map (* we add it in each args_res *) + (fun res2 -> + combine_fun res1 res2 + ) + res2.result ) res1.result in (* and then we flatten the map *) @@ -127,18 +127,18 @@ let rec change_vars_in_binder mapping = function | (bt,t)::l -> let new_mapping = Id.Set.fold Id.Map.remove (ids_of_binder bt) mapping in (bt,change_vars mapping t):: - (if Id.Map.is_empty new_mapping - then l - else change_vars_in_binder new_mapping l - ) + (if Id.Map.is_empty new_mapping + then l + else change_vars_in_binder new_mapping l + ) let rec replace_var_by_term_in_binder x_id term = function | [] -> [] | (bt,t)::l -> (bt,replace_var_by_term x_id term t):: - if Id.Set.mem x_id (ids_of_binder bt) - then l - else replace_var_by_term_in_binder x_id term l + if Id.Set.mem x_id (ids_of_binder bt) + then l + else replace_var_by_term_in_binder x_id term l let add_bt_names bt = Id.Set.union (ids_of_binder bt) @@ -152,66 +152,66 @@ let apply_args ctxt body args = let next_name_away (na:Name.t) (mapping: Id.t Id.Map.t) (avoid: Id.Set.t) = match na with | Name id when Id.Set.mem id avoid -> - let new_id = Namegen.next_ident_away id avoid in - Name new_id,Id.Map.add id new_id mapping,Id.Set.add new_id avoid + let new_id = Namegen.next_ident_away id avoid in + Name new_id,Id.Map.add id new_id mapping,Id.Set.add new_id avoid | _ -> na,mapping,avoid in let next_bt_away bt (avoid:Id.Set.t) = match bt with | LetIn na -> - let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in - LetIn new_na,mapping,new_avoid + let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in + LetIn new_na,mapping,new_avoid | Prod na -> - let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in - Prod new_na,mapping,new_avoid + let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in + Prod new_na,mapping,new_avoid | Lambda na -> - let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in - Lambda new_na,mapping,new_avoid + let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in + Lambda new_na,mapping,new_avoid in let rec do_apply avoid ctxt body args = match ctxt,args with | _,[] -> (* No more args *) - (ctxt,body) + (ctxt,body) | [],_ -> (* no more fun *) - let f,args' = glob_decompose_app body in - (ctxt,mkGApp(f,args'@args)) + let f,args' = glob_decompose_app body in + (ctxt,mkGApp(f,args'@args)) | (Lambda Anonymous,t)::ctxt',arg::args' -> - do_apply avoid ctxt' body args' + do_apply avoid ctxt' body args' | (Lambda (Name id),t)::ctxt',arg::args' -> - let new_avoid,new_ctxt',new_body,new_id = - if need_convert_id avoid id - then - let new_avoid = Id.Set.add id avoid in - let new_id = Namegen.next_ident_away id new_avoid in - let new_avoid' = Id.Set.add new_id new_avoid in - let mapping = Id.Map.add id new_id Id.Map.empty in - let new_ctxt' = change_vars_in_binder mapping ctxt' in - let new_body = change_vars mapping body in - new_avoid',new_ctxt',new_body,new_id - else - Id.Set.add id avoid,ctxt',body,id - in - let new_body = replace_var_by_term new_id arg new_body in - let new_ctxt' = replace_var_by_term_in_binder new_id arg new_ctxt' in - do_apply avoid new_ctxt' new_body args' + let new_avoid,new_ctxt',new_body,new_id = + if need_convert_id avoid id + then + let new_avoid = Id.Set.add id avoid in + let new_id = Namegen.next_ident_away id new_avoid in + let new_avoid' = Id.Set.add new_id new_avoid in + let mapping = Id.Map.add id new_id Id.Map.empty in + let new_ctxt' = change_vars_in_binder mapping ctxt' in + let new_body = change_vars mapping body in + new_avoid',new_ctxt',new_body,new_id + else + Id.Set.add id avoid,ctxt',body,id + in + let new_body = replace_var_by_term new_id arg new_body in + let new_ctxt' = replace_var_by_term_in_binder new_id arg new_ctxt' in + do_apply avoid new_ctxt' new_body args' | (bt,t)::ctxt',_ -> - let new_avoid,new_ctxt',new_body,new_bt = - let new_avoid = add_bt_names bt avoid in - if need_convert avoid bt - then - let new_bt,mapping,new_avoid = next_bt_away bt new_avoid in - ( - new_avoid, - change_vars_in_binder mapping ctxt', - change_vars mapping body, - new_bt - ) - else new_avoid,ctxt',body,bt - in - let new_ctxt',new_body = - do_apply new_avoid new_ctxt' new_body args - in - (new_bt,t)::new_ctxt',new_body + let new_avoid,new_ctxt',new_body,new_bt = + let new_avoid = add_bt_names bt avoid in + if need_convert avoid bt + then + let new_bt,mapping,new_avoid = next_bt_away bt new_avoid in + ( + new_avoid, + change_vars_in_binder mapping ctxt', + change_vars mapping body, + new_bt + ) + else new_avoid,ctxt',body,bt + in + let new_ctxt',new_body = + do_apply new_avoid new_ctxt' new_body args + in + (new_bt,t)::new_ctxt',new_body in do_apply Id.Set.empty ctxt body args @@ -230,14 +230,14 @@ let combine_lam n t b = { context = []; value = mkGLambda(n, compose_glob_context t.context t.value, - compose_glob_context b.context b.value ) + compose_glob_context b.context b.value ) } let combine_prod2 n t b = { context = []; value = mkGProd(n, compose_glob_context t.context t.value, - compose_glob_context b.context b.value ) + compose_glob_context b.context b.value ) } let combine_prod n t b = @@ -251,8 +251,8 @@ let mk_result ctxt value avoid = { result = [{context = ctxt; - value = value}] - ; + value = value}] + ; to_avoid = avoid } (************************************************* @@ -298,8 +298,8 @@ let make_discr_match_brl i = let make_discr_match brl = fun el i -> mkGCases(None, - make_discr_match_el el, - make_discr_match_brl i brl) + make_discr_match_el el, + make_discr_match_brl i brl) (**********************************************************************) (* functions used to build case expression from lettuple and if ones *) @@ -310,27 +310,27 @@ let build_constructors_of_type ind' argl = let (mib,ind) = Inductive.lookup_mind_specif (Global.env()) ind' in let npar = mib.Declarations.mind_nparams in Array.mapi (fun i _ -> - let construct = ind',i+1 in + let construct = ind',i+1 in let constructref = GlobRef.ConstructRef(construct) in - let _implicit_positions_of_cst = - Impargs.implicits_of_global constructref - in - let cst_narg = + let _implicit_positions_of_cst = + Impargs.implicits_of_global constructref + in + let cst_narg = Inductiveops.constructor_nallargs - (Global.env ()) - construct - in - let argl = + (Global.env ()) + construct + in + let argl = if List.is_empty argl then List.make cst_narg (mkGHole ()) else List.make npar (mkGHole ()) @ argl - in - let pat_as_term = + in + let pat_as_term = mkGApp(mkGRef (GlobRef.ConstructRef(ind',i+1)),argl) - in + in cases_pattern_of_glob_constr (Global.env()) Anonymous pat_as_term - ) + ) ind.Declarations.mind_consnames (******************) @@ -359,20 +359,20 @@ let add_pat_variables sigma pat typ env : Environ.env = match DAst.get pat with | PatVar na -> Environ.push_rel (RelDecl.LocalAssum (make_annot na Sorts.Relevant,typ)) env | PatCstr(c,patl,na) -> - let Inductiveops.IndType(indf,indargs) = - try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) - with Not_found -> assert false - in - let constructors = Inductiveops.get_constructors env indf in - let constructor : Inductiveops.constructor_summary = List.find (fun cs -> eq_constructor c (fst cs.Inductiveops.cs_cstr)) (Array.to_list constructors) in - let cs_args_types :types list = List.map RelDecl.get_type constructor.Inductiveops.cs_args in - List.fold_left2 add_pat_variables env patl (List.rev cs_args_types) + let Inductiveops.IndType(indf,indargs) = + try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) + with Not_found -> assert false + in + let constructors = Inductiveops.get_constructors env indf in + let constructor : Inductiveops.constructor_summary = List.find (fun cs -> eq_constructor c (fst cs.Inductiveops.cs_cstr)) (Array.to_list constructors) in + let cs_args_types :types list = List.map RelDecl.get_type constructor.Inductiveops.cs_args in + List.fold_left2 add_pat_variables env patl (List.rev cs_args_types) in let new_env = add_pat_variables env pat typ in let res = fst ( Context.Rel.fold_outside - (fun decl (env,ctxt) -> + (fun decl (env,ctxt) -> let open Context.Rel.Declaration in match decl with | LocalAssum ({binder_name=Anonymous},_) | LocalDef ({binder_name=Anonymous},_,_) -> assert false @@ -398,8 +398,8 @@ let add_pat_variables sigma pat typ env : Environ.env = let open Context.Named.Declaration in (Environ.push_named (LocalDef (na,new_v,new_t)) env,mkVar id::ctxt) ) - (Environ.rel_context new_env) - ~init:(env,[]) + (Environ.rel_context new_env) + ~init:(env,[]) ) in observe (str "new var env := " ++ Printer.pr_named_context_of res (Evd.from_env env)); @@ -411,16 +411,16 @@ let add_pat_variables sigma pat typ env : Environ.env = let rec pattern_to_term_and_type env typ = DAst.with_val (function | PatVar Anonymous -> assert false | PatVar (Name id) -> - mkGVar id + mkGVar id | PatCstr(constr,patternl,_) -> let cst_narg = Inductiveops.constructor_nallargs - (Global.env ()) - constr + (Global.env ()) + constr in let Inductiveops.IndType(indf,indargs) = - try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) - with Not_found -> assert false + try Inductiveops.find_rectype env (Evd.from_env env) (EConstr.of_constr typ) + with Not_found -> assert false in let constructors = Inductiveops.get_constructors env indf in let constructor = List.find (fun cs -> eq_constructor (fst cs.Inductiveops.cs_cstr) constr) (Array.to_list constructors) in @@ -428,18 +428,18 @@ let rec pattern_to_term_and_type env typ = DAst.with_val (function let _,cstl = Inductiveops.dest_ind_family indf in let csta = Array.of_list cstl in let implicit_args = - Array.to_list - (Array.init - (cst_narg - List.length patternl) - (fun i -> Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) (EConstr.of_constr csta.(i))) - ) + Array.to_list + (Array.init + (cst_narg - List.length patternl) + (fun i -> Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) (EConstr.of_constr csta.(i))) + ) in let patl_as_term = - List.map2 (pattern_to_term_and_type env) (List.rev cs_args_types) patternl + List.map2 (pattern_to_term_and_type env) (List.rev cs_args_types) patternl in mkGApp(mkGRef(GlobRef.ConstructRef constr), - implicit_args@patl_as_term - ) + implicit_args@patl_as_term + ) ) (* [build_entry_lc funnames avoid rt] construct the list (in fact a build_entry_return) @@ -479,220 +479,220 @@ let rec build_entry_lc env sigma funnames avoid rt : glob_constr build_entry_ret let open CAst in match DAst.get rt with | GRef _ | GVar _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ | GFloat _ -> - (* do nothing (except changing type of course) *) - mk_result [] rt avoid + (* do nothing (except changing type of course) *) + mk_result [] rt avoid | GApp(_,_) -> - let f,args = glob_decompose_app rt in - let args_res : (glob_constr list) build_entry_return = - List.fold_right (* create the arguments lists of constructors and combine them *) - (fun arg ctxt_argsl -> + let f,args = glob_decompose_app rt in + let args_res : (glob_constr list) build_entry_return = + List.fold_right (* create the arguments lists of constructors and combine them *) + (fun arg ctxt_argsl -> let arg_res = build_entry_lc env sigma funnames ctxt_argsl.to_avoid arg in - combine_results combine_args arg_res ctxt_argsl - ) - args - (mk_result [] [] avoid) - in - begin - match DAst.get f with - | GLambda _ -> - let rec aux t l = - match l with - | [] -> t - | u::l -> DAst.make @@ - match DAst.get t with - | GLambda(na,_,nat,b) -> - GLetIn(na,u,None,aux b l) - | _ -> - GApp(t,l) - in + combine_results combine_args arg_res ctxt_argsl + ) + args + (mk_result [] [] avoid) + in + begin + match DAst.get f with + | GLambda _ -> + let rec aux t l = + match l with + | [] -> t + | u::l -> DAst.make @@ + match DAst.get t with + | GLambda(na,_,nat,b) -> + GLetIn(na,u,None,aux b l) + | _ -> + GApp(t,l) + in build_entry_lc env sigma funnames avoid (aux f args) - | GVar id when Id.Set.mem id funnames -> - (* if we have [f t1 ... tn] with [f]$\in$[fnames] - then we create a fresh variable [res], - add [res] and its "value" (i.e. [res v1 ... vn]) to each - pseudo constructor build for the arguments (i.e. a pseudo context [ctxt] and - a pseudo value "v1 ... vn". - The "value" of this branch is then simply [res] - *) - let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in + | GVar id when Id.Set.mem id funnames -> + (* if we have [f t1 ... tn] with [f]$\in$[fnames] + then we create a fresh variable [res], + add [res] and its "value" (i.e. [res v1 ... vn]) to each + pseudo constructor build for the arguments (i.e. a pseudo context [ctxt] and + a pseudo value "v1 ... vn". + The "value" of this branch is then simply [res] + *) + let rt_as_constr,ctx = Pretyping.understand env (Evd.from_env env) rt in let rt_typ = Typing.unsafe_type_of env (Evd.from_env env) rt_as_constr in - let res_raw_type = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) rt_typ in - let res = fresh_id args_res.to_avoid "_res" in - let new_avoid = res::args_res.to_avoid in - let res_rt = mkGVar res in - let new_result = - List.map - (fun arg_res -> - let new_hyps = - [Prod (Name res),res_raw_type; - Prod Anonymous,mkGApp(res_rt,(mkGVar id)::arg_res.value)] - in - {context = arg_res.context@new_hyps; value = res_rt } - ) - args_res.result - in - { result = new_result; to_avoid = new_avoid } - | GVar _ | GEvar _ | GPatVar _ | GHole _ | GSort _ | GRef _ -> - (* if have [g t1 ... tn] with [g] not appearing in [funnames] - then - foreach [ctxt,v1 ... vn] in [args_res] we return - [ctxt, g v1 .... vn] - *) - { - args_res with - result = - List.map - (fun args_res -> - {args_res with value = mkGApp(f,args_res.value)}) - args_res.result - } - | GApp _ -> assert false (* we have collected all the app in [glob_decompose_app] *) - | GLetIn(n,v,t,b) -> - (* if we have [(let x := v in b) t1 ... tn] , - we discard our work and compute the list of constructor for - [let x = v in (b t1 ... tn)] up to alpha conversion - *) - let new_n,new_b,new_avoid = - match n with - | Name id when List.exists (is_free_in id) args -> - (* need to alpha-convert the name *) - let new_id = Namegen.next_ident_away id (Id.Set.of_list avoid) in - let new_avoid = id:: avoid in - let new_b = - replace_var_by_term - id - (DAst.make @@ GVar id) - b - in - (Name new_id,new_b,new_avoid) - | _ -> n,b,avoid - in - build_entry_lc - env + let res_raw_type = Detyping.detype Detyping.Now false Id.Set.empty env (Evd.from_env env) rt_typ in + let res = fresh_id args_res.to_avoid "_res" in + let new_avoid = res::args_res.to_avoid in + let res_rt = mkGVar res in + let new_result = + List.map + (fun arg_res -> + let new_hyps = + [Prod (Name res),res_raw_type; + Prod Anonymous,mkGApp(res_rt,(mkGVar id)::arg_res.value)] + in + {context = arg_res.context@new_hyps; value = res_rt } + ) + args_res.result + in + { result = new_result; to_avoid = new_avoid } + | GVar _ | GEvar _ | GPatVar _ | GHole _ | GSort _ | GRef _ -> + (* if have [g t1 ... tn] with [g] not appearing in [funnames] + then + foreach [ctxt,v1 ... vn] in [args_res] we return + [ctxt, g v1 .... vn] + *) + { + args_res with + result = + List.map + (fun args_res -> + {args_res with value = mkGApp(f,args_res.value)}) + args_res.result + } + | GApp _ -> assert false (* we have collected all the app in [glob_decompose_app] *) + | GLetIn(n,v,t,b) -> + (* if we have [(let x := v in b) t1 ... tn] , + we discard our work and compute the list of constructor for + [let x = v in (b t1 ... tn)] up to alpha conversion + *) + let new_n,new_b,new_avoid = + match n with + | Name id when List.exists (is_free_in id) args -> + (* need to alpha-convert the name *) + let new_id = Namegen.next_ident_away id (Id.Set.of_list avoid) in + let new_avoid = id:: avoid in + let new_b = + replace_var_by_term + id + (DAst.make @@ GVar id) + b + in + (Name new_id,new_b,new_avoid) + | _ -> n,b,avoid + in + build_entry_lc + env sigma funnames - avoid - (mkGLetIn(new_n,v,t,mkGApp(new_b,args))) - | GCases _ | GIf _ | GLetTuple _ -> - (* we have [(match e1, ...., en with ..... end) t1 tn] - we first compute the result from the case and - then combine each of them with each of args one - *) + avoid + (mkGLetIn(new_n,v,t,mkGApp(new_b,args))) + | GCases _ | GIf _ | GLetTuple _ -> + (* we have [(match e1, ...., en with ..... end) t1 tn] + we first compute the result from the case and + then combine each of them with each of args one + *) let f_res = build_entry_lc env sigma funnames args_res.to_avoid f in - combine_results combine_app f_res args_res - | GCast(b,_) -> - (* for an applied cast we just trash the cast part - and restart the work. + combine_results combine_app f_res args_res + | GCast(b,_) -> + (* for an applied cast we just trash the cast part + and restart the work. - WARNING: We need to restart since [b] itself should be an application term - *) + WARNING: We need to restart since [b] itself should be an application term + *) build_entry_lc env sigma funnames avoid (mkGApp(b,args)) - | GRec _ -> user_err Pp.(str "Not handled GRec") - | GProd _ -> user_err Pp.(str "Cannot apply a type") + | GRec _ -> user_err Pp.(str "Not handled GRec") + | GProd _ -> user_err Pp.(str "Cannot apply a type") | GInt _ -> user_err Pp.(str "Cannot apply an integer") | GFloat _ -> user_err Pp.(str "Cannot apply a float") - end (* end of the application treatement *) + end (* end of the application treatement *) | GLambda(n,_,t,b) -> - (* we first compute the list of constructor - corresponding to the body of the function, - then the one corresponding to the type - and combine the two result - *) + (* we first compute the list of constructor + corresponding to the body of the function, + then the one corresponding to the type + and combine the two result + *) let t_res = build_entry_lc env sigma funnames avoid t in - let new_n = - match n with - | Name _ -> n - | Anonymous -> Name (Indfun_common.fresh_id [] "_x") - in - let new_env = raw_push_named (new_n,None,t) env in + let new_n = + match n with + | Name _ -> n + | Anonymous -> Name (Indfun_common.fresh_id [] "_x") + in + let new_env = raw_push_named (new_n,None,t) env in let b_res = build_entry_lc new_env sigma funnames avoid b in - combine_results (combine_lam new_n) t_res b_res + combine_results (combine_lam new_n) t_res b_res | GProd(n,_,t,b) -> - (* we first compute the list of constructor - corresponding to the body of the function, - then the one corresponding to the type - and combine the two result - *) + (* we first compute the list of constructor + corresponding to the body of the function, + then the one corresponding to the type + and combine the two result + *) let t_res = build_entry_lc env sigma funnames avoid t in - let new_env = raw_push_named (n,None,t) env in + let new_env = raw_push_named (n,None,t) env in let b_res = build_entry_lc new_env sigma funnames avoid b in if List.length t_res.result = 1 && List.length b_res.result = 1 then combine_results (combine_prod2 n) t_res b_res else combine_results (combine_prod n) t_res b_res | GLetIn(n,v,typ,b) -> - (* we first compute the list of constructor - corresponding to the body of the function, - then the one corresponding to the value [t] - and combine the two result - *) + (* we first compute the list of constructor + corresponding to the body of the function, + then the one corresponding to the value [t] + and combine the two result + *) let v = match typ with None -> v | Some t -> DAst.make ?loc:rt.loc @@ GCast (v,CastConv t) in let v_res = build_entry_lc env sigma funnames avoid v in - let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in + let v_as_constr,ctx = Pretyping.understand env (Evd.from_env env) v in let v_type = Typing.unsafe_type_of env (Evd.from_env env) v_as_constr in let v_r = Sorts.Relevant in (* TODO relevance *) - let new_env = - match n with - Anonymous -> env + let new_env = + match n with + Anonymous -> env | Name id -> EConstr.push_named (NamedDecl.LocalDef (make_annot id v_r,v_as_constr,v_type)) env in let b_res = build_entry_lc new_env sigma funnames avoid b in - combine_results (combine_letin n) v_res b_res + combine_results (combine_letin n) v_res b_res | GCases(_,_,el,brl) -> - (* we create the discrimination function - and treat the case itself - *) - let make_discr = make_discr_match brl in + (* we create the discrimination function + and treat the case itself + *) + let make_discr = make_discr_match brl in build_entry_lc_from_case env sigma funnames make_discr el brl avoid | GIf(b,(na,e_option),lhs,rhs) -> - let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in + let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in - let (ind,_) = - try Inductiveops.find_inductive env (Evd.from_env env) b_typ - with Not_found -> - user_err (str "Cannot find the inductive associated to " ++ + let (ind,_) = + try Inductiveops.find_inductive env (Evd.from_env env) b_typ + with Not_found -> + user_err (str "Cannot find the inductive associated to " ++ Printer.pr_glob_constr_env env b ++ str " in " ++ Printer.pr_glob_constr_env env rt ++ str ". try again with a cast") - in - let case_pats = build_constructors_of_type (fst ind) [] in - assert (Int.equal (Array.length case_pats) 2); - let brl = - List.map_i + in + let case_pats = build_constructors_of_type (fst ind) [] in + assert (Int.equal (Array.length case_pats) 2); + let brl = + List.map_i (fun i x -> CAst.make ([],[case_pats.(i)],x)) - 0 - [lhs;rhs] - in - let match_expr = - mkGCases(None,[(b,(Anonymous,None))],brl) - in - (* Pp.msgnl (str "new case := " ++ Printer.pr_glob_constr match_expr); *) + 0 + [lhs;rhs] + in + let match_expr = + mkGCases(None,[(b,(Anonymous,None))],brl) + in + (* Pp.msgnl (str "new case := " ++ Printer.pr_glob_constr match_expr); *) build_entry_lc env sigma funnames avoid match_expr | GLetTuple(nal,_,b,e) -> - begin - let nal_as_glob_constr = - List.map - (function - Name id -> mkGVar id - | Anonymous -> mkGHole () - ) - nal - in - let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in + begin + let nal_as_glob_constr = + List.map + (function + Name id -> mkGVar id + | Anonymous -> mkGHole () + ) + nal + in + let b_as_constr,ctx = Pretyping.understand env (Evd.from_env env) b in let b_typ = Typing.unsafe_type_of env (Evd.from_env env) b_as_constr in - let (ind,_) = - try Inductiveops.find_inductive env (Evd.from_env env) b_typ - with Not_found -> - user_err (str "Cannot find the inductive associated to " ++ + let (ind,_) = + try Inductiveops.find_inductive env (Evd.from_env env) b_typ + with Not_found -> + user_err (str "Cannot find the inductive associated to " ++ Printer.pr_glob_constr_env env b ++ str " in " ++ Printer.pr_glob_constr_env env rt ++ str ". try again with a cast") - in - let case_pats = build_constructors_of_type (fst ind) nal_as_glob_constr in - assert (Int.equal (Array.length case_pats) 1); + in + let case_pats = build_constructors_of_type (fst ind) nal_as_glob_constr in + assert (Int.equal (Array.length case_pats) 1); let br = CAst.make ([],[case_pats.(0)],e) in - let match_expr = mkGCases(None,[b,(Anonymous,None)],[br]) in + let match_expr = mkGCases(None,[b,(Anonymous,None)],[br]) in build_entry_lc env sigma funnames avoid match_expr - end + end | GRec _ -> user_err Pp.(str "Not handled GRec") | GCast(b,_) -> build_entry_lc env sigma funnames avoid b @@ -703,177 +703,177 @@ and build_entry_lc_from_case env sigma funname make_discr match el with | [] -> assert false (* this case correspond to match <nothing> with .... !*) | el -> - (* this case correspond to - match el with brl end - we first compute the list of lists corresponding to [el] and - combine them . - Then for each element of the combinations, - we compute the result we compute one list per branch in [brl] and - finally we just concatenate those list - *) - let case_resl = - List.fold_right - (fun (case_arg,_) ctxt_argsl -> + (* this case correspond to + match el with brl end + we first compute the list of lists corresponding to [el] and + combine them . + Then for each element of the combinations, + we compute the result we compute one list per branch in [brl] and + finally we just concatenate those list + *) + let case_resl = + List.fold_right + (fun (case_arg,_) ctxt_argsl -> let arg_res = build_entry_lc env sigma funname ctxt_argsl.to_avoid case_arg in - combine_results combine_args arg_res ctxt_argsl - ) - el - (mk_result [] [] avoid) - in - let types = - List.map (fun (case_arg,_) -> - let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in + combine_results combine_args arg_res ctxt_argsl + ) + el + (mk_result [] [] avoid) + in + let types = + List.map (fun (case_arg,_) -> + let case_arg_as_constr,ctx = Pretyping.understand env (Evd.from_env env) case_arg in EConstr.Unsafe.to_constr (Typing.unsafe_type_of env (Evd.from_env env) case_arg_as_constr) - ) el - in - (****** The next works only if the match is not dependent ****) - let results = - List.map - (fun ca -> - let res = build_entry_lc_from_case_term + ) el + in + (****** The next works only if the match is not dependent ****) + let results = + List.map + (fun ca -> + let res = build_entry_lc_from_case_term env sigma types - funname (make_discr) - [] brl - case_resl.to_avoid - ca - in - res - ) - case_resl.result - in - { - result = List.concat (List.map (fun r -> r.result) results); - to_avoid = - List.fold_left (fun acc r -> List.union Id.equal acc r.to_avoid) + funname (make_discr) + [] brl + case_resl.to_avoid + ca + in + res + ) + case_resl.result + in + { + result = List.concat (List.map (fun r -> r.result) results); + to_avoid = + List.fold_left (fun acc r -> List.union Id.equal acc r.to_avoid) [] results - } + } and build_entry_lc_from_case_term env sigma types funname make_discr patterns_to_prevent brl avoid matched_expr = match brl with | [] -> (* computed_branches *) {result = [];to_avoid = avoid} | br::brl' -> - (* alpha conversion to prevent name clashes *) + (* alpha conversion to prevent name clashes *) let {CAst.v=(idl,patl,return)} = alpha_br avoid br in - let new_avoid = idl@avoid in (* for now we can no more use idl as an identifier *) - (* building a list of precondition stating that we are not in this branch - (will be used in the following recursive calls) - *) + let new_avoid = idl@avoid in (* for now we can no more use idl as an identifier *) + (* building a list of precondition stating that we are not in this branch + (will be used in the following recursive calls) + *) let new_env = List.fold_right2 (add_pat_variables sigma) patl types env in - let not_those_patterns : (Id.t list -> glob_constr -> glob_constr) list = - List.map2 - (fun pat typ -> - fun avoid pat'_as_term -> - let renamed_pat,_,_ = alpha_pat avoid pat in - let pat_ids = get_pattern_id renamed_pat in + let not_those_patterns : (Id.t list -> glob_constr -> glob_constr) list = + List.map2 + (fun pat typ -> + fun avoid pat'_as_term -> + let renamed_pat,_,_ = alpha_pat avoid pat in + let pat_ids = get_pattern_id renamed_pat in let env_with_pat_ids = add_pat_variables sigma pat typ new_env in - List.fold_right - (fun id acc -> - let typ_of_id = - Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (EConstr.mkVar id) - in - let raw_typ_of_id = - Detyping.detype Detyping.Now false Id.Set.empty - env_with_pat_ids (Evd.from_env env) typ_of_id - in - mkGProd (Name id,raw_typ_of_id,acc)) - pat_ids - (glob_make_neq pat'_as_term (pattern_to_term renamed_pat)) - ) - patl - types - in - (* Checking if we can be in this branch - (will be used in the following recursive calls) - *) - let unify_with_those_patterns : (cases_pattern -> bool*bool) list = - List.map - (fun pat pat' -> are_unifiable pat pat',eq_cases_pattern pat pat') - patl - in - (* - we first compute the other branch result (in ordrer to keep the order of the matching - as much as possible) - *) - let brl'_res = - build_entry_lc_from_case_term - env + List.fold_right + (fun id acc -> + let typ_of_id = + Typing.unsafe_type_of env_with_pat_ids (Evd.from_env env) (EConstr.mkVar id) + in + let raw_typ_of_id = + Detyping.detype Detyping.Now false Id.Set.empty + env_with_pat_ids (Evd.from_env env) typ_of_id + in + mkGProd (Name id,raw_typ_of_id,acc)) + pat_ids + (glob_make_neq pat'_as_term (pattern_to_term renamed_pat)) + ) + patl + types + in + (* Checking if we can be in this branch + (will be used in the following recursive calls) + *) + let unify_with_those_patterns : (cases_pattern -> bool*bool) list = + List.map + (fun pat pat' -> are_unifiable pat pat',eq_cases_pattern pat pat') + patl + in + (* + we first compute the other branch result (in ordrer to keep the order of the matching + as much as possible) + *) + let brl'_res = + build_entry_lc_from_case_term + env sigma - types - funname - make_discr - ((unify_with_those_patterns,not_those_patterns)::patterns_to_prevent) - brl' - avoid - matched_expr - in - (* We now create the precondition of this branch i.e. - 1- the list of variable appearing in the different patterns of this branch and - the list of equation stating than el = patl (List.flatten ...) - 2- If there exists a previous branch which pattern unify with the one of this branch + types + funname + make_discr + ((unify_with_those_patterns,not_those_patterns)::patterns_to_prevent) + brl' + avoid + matched_expr + in + (* We now create the precondition of this branch i.e. + 1- the list of variable appearing in the different patterns of this branch and + the list of equation stating than el = patl (List.flatten ...) + 2- If there exists a previous branch which pattern unify with the one of this branch then a discrimination precond stating that we are not in a previous branch (if List.exists ...) - *) - let those_pattern_preconds = - (List.flatten - ( - List.map3 - (fun pat e typ_as_constr -> - let this_pat_ids = ids_of_pat pat in - let typ_as_constr = EConstr.of_constr typ_as_constr in - let typ = Detyping.detype Detyping.Now false Id.Set.empty new_env (Evd.from_env env) typ_as_constr in - let pat_as_term = pattern_to_term pat in - (* removing trivial holes *) - let pat_as_term = solve_trivial_holes pat_as_term e in + *) + let those_pattern_preconds = + (List.flatten + ( + List.map3 + (fun pat e typ_as_constr -> + let this_pat_ids = ids_of_pat pat in + let typ_as_constr = EConstr.of_constr typ_as_constr in + let typ = Detyping.detype Detyping.Now false Id.Set.empty new_env (Evd.from_env env) typ_as_constr in + let pat_as_term = pattern_to_term pat in + (* removing trivial holes *) + let pat_as_term = solve_trivial_holes pat_as_term e in (* observe (str "those_pattern_preconds" ++ spc () ++ *) (* str "pat" ++ spc () ++ pr_glob_constr pat_as_term ++ spc ()++ *) (* str "e" ++ spc () ++ pr_glob_constr e ++spc ()++ *) (* str "typ_as_constr" ++ spc () ++ pr_lconstr typ_as_constr); *) - List.fold_right - (fun id acc -> - if Id.Set.mem id this_pat_ids - then (Prod (Name id), - let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (EConstr.mkVar id) in - let raw_typ_of_id = - Detyping.detype Detyping.Now false Id.Set.empty new_env (Evd.from_env env) typ_of_id - in - raw_typ_of_id - )::acc - else acc - ) - idl - [(Prod Anonymous,glob_make_eq ~typ pat_as_term e)] - ) - patl - matched_expr.value - types - ) - ) - @ - (if List.exists (function (unifl,_) -> - let (unif,_) = - List.split (List.map2 (fun x y -> x y) unifl patl) - in - List.for_all (fun x -> x) unif) patterns_to_prevent - then - let i = List.length patterns_to_prevent in - let pats_as_constr = List.map2 (pattern_to_term_and_type new_env) types patl in - [(Prod Anonymous,make_discr pats_as_constr i )] - else - [] - ) - in - (* We compute the result of the value returned by the branch*) + List.fold_right + (fun id acc -> + if Id.Set.mem id this_pat_ids + then (Prod (Name id), + let typ_of_id = Typing.unsafe_type_of new_env (Evd.from_env env) (EConstr.mkVar id) in + let raw_typ_of_id = + Detyping.detype Detyping.Now false Id.Set.empty new_env (Evd.from_env env) typ_of_id + in + raw_typ_of_id + )::acc + else acc + ) + idl + [(Prod Anonymous,glob_make_eq ~typ pat_as_term e)] + ) + patl + matched_expr.value + types + ) + ) + @ + (if List.exists (function (unifl,_) -> + let (unif,_) = + List.split (List.map2 (fun x y -> x y) unifl patl) + in + List.for_all (fun x -> x) unif) patterns_to_prevent + then + let i = List.length patterns_to_prevent in + let pats_as_constr = List.map2 (pattern_to_term_and_type new_env) types patl in + [(Prod Anonymous,make_discr pats_as_constr i )] + else + [] + ) + in + (* We compute the result of the value returned by the branch*) let return_res = build_entry_lc new_env sigma funname new_avoid return in - (* and combine it with the preconds computed for this branch *) - let this_branch_res = - List.map - (fun res -> - { context = matched_expr.context@those_pattern_preconds@res.context ; - value = res.value} - ) - return_res.result - in - { brl'_res with result = this_branch_res@brl'_res.result } + (* and combine it with the preconds computed for this branch *) + let this_branch_res = + List.map + (fun res -> + { context = matched_expr.context@those_pattern_preconds@res.context ; + value = res.value} + ) + return_res.result + in + { brl'_res with result = this_branch_res@brl'_res.result } let is_res r = match DAst.get r with @@ -891,8 +891,8 @@ let is_gvar c = match DAst.get c with | GVar id -> true | _ -> false -let same_raw_term rt1 rt2 = - match DAst.get rt1, DAst.get rt2 with +let same_raw_term rt1 rt2 = + match DAst.get rt1, DAst.get rt2 with | GRef(r1,_), GRef (r2,_) -> GlobRef.equal r1 r2 | GHole _, GHole _ -> true | _ -> false @@ -927,288 +927,288 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt = let open CAst in match DAst.get rt with | GProd(n,k,t,b) -> - let not_free_in_t id = not (is_free_in id t) in - let new_crossed_types = t::crossed_types in - begin - match DAst.get t with - | GApp(res_rt ,args') when is_res res_rt -> - begin + let not_free_in_t id = not (is_free_in id t) in + let new_crossed_types = t::crossed_types in + begin + match DAst.get t with + | GApp(res_rt ,args') when is_res res_rt -> + begin let arg = List.hd args' in - match DAst.get arg with - | GVar this_relname -> - (*i The next call to mk_rel_id is - valid since we are constructing the graph - Ensures by: obvious - i*) - - let new_t = - mkGApp(mkGVar(mk_rel_id this_relname),List.tl args'@[res_rt]) - in + match DAst.get arg with + | GVar this_relname -> + (*i The next call to mk_rel_id is + valid since we are constructing the graph + Ensures by: obvious + i*) + + let new_t = + mkGApp(mkGVar(mk_rel_id this_relname),List.tl args'@[res_rt]) + in let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = - rebuild_cons new_env - nb_args relname - args new_crossed_types - (depth + 1) b - in - mkGProd(n,new_t,new_b), - Id.Set.filter not_free_in_t id_to_exclude - | _ -> (* the first args is the name of the function! *) - assert false - end - | GApp(eq_as_ref,[ty; id ;rt]) + rebuild_cons new_env + nb_args relname + args new_crossed_types + (depth + 1) b + in + mkGProd(n,new_t,new_b), + Id.Set.filter not_free_in_t id_to_exclude + | _ -> (* the first args is the name of the function! *) + assert false + end + | GApp(eq_as_ref,[ty; id ;rt]) when is_gvar id && is_gr eq_as_ref Coqlib.(lib_ref "core.eq.type") && n == Anonymous - -> + -> let loc1 = rt.CAst.loc in let loc2 = eq_as_ref.CAst.loc in let loc3 = id.CAst.loc in let id = match DAst.get id with GVar id -> id | _ -> assert false in - begin - try + begin + try observe (str "computing new type for eq : " ++ pr_glob_constr_env env rt); - let t' = - try fst (Pretyping.understand env (Evd.from_env env) t)(*FIXME*) + let t' = + try fst (Pretyping.understand env (Evd.from_env env) t)(*FIXME*) with e when CErrors.noncritical e -> raise Continue - in - let is_in_b = is_free_in id b in - let _keep_eq = - not (List.exists (is_free_in id) args) || is_in_b || - List.exists (is_free_in id) crossed_types - in - let new_args = List.map (replace_var_by_term id rt) args in - let subst_b = - if is_in_b then b else replace_var_by_term id rt b + in + let is_in_b = is_free_in id b in + let _keep_eq = + not (List.exists (is_free_in id) args) || is_in_b || + List.exists (is_free_in id) crossed_types + in + let new_args = List.map (replace_var_by_term id rt) args in + let subst_b = + if is_in_b then b else replace_var_by_term id rt b in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = - rebuild_cons - new_env - nb_args relname - new_args new_crossed_types - (depth + 1) subst_b - in - mkGProd(n,t,new_b),id_to_exclude - with Continue -> + rebuild_cons + new_env + nb_args relname + new_args new_crossed_types + (depth + 1) subst_b + in + mkGProd(n,t,new_b),id_to_exclude + with Continue -> let jmeq = GlobRef.IndRef (fst (EConstr.destInd Evd.empty (jmeq ()))) in - let ty',ctx = Pretyping.understand env (Evd.from_env env) ty in + let ty',ctx = Pretyping.understand env (Evd.from_env env) ty in let ind,args' = Inductiveops.find_inductive env Evd.(from_env env) ty' in - let mib,_ = Global.lookup_inductive (fst ind) in - let nparam = mib.Declarations.mind_nparams in - let params,arg' = - ((Util.List.chop nparam args')) - in - let rt_typ = DAst.make @@ + let mib,_ = Global.lookup_inductive (fst ind) in + let nparam = mib.Declarations.mind_nparams in + let params,arg' = + ((Util.List.chop nparam args')) + in + let rt_typ = DAst.make @@ GApp(DAst.make @@ GRef (GlobRef.IndRef (fst ind),None), - (List.map - (fun p -> Detyping.detype Detyping.Now false Id.Set.empty - env (Evd.from_env env) - (EConstr.of_constr p)) params)@(Array.to_list - (Array.make - (List.length args' - nparam) - (mkGHole ())))) - in - let eq' = - DAst.make ?loc:loc1 @@ GApp(DAst.make ?loc:loc2 @@GRef(jmeq,None),[ty;DAst.make ?loc:loc3 @@ GVar id;rt_typ;rt]) - in + (List.map + (fun p -> Detyping.detype Detyping.Now false Id.Set.empty + env (Evd.from_env env) + (EConstr.of_constr p)) params)@(Array.to_list + (Array.make + (List.length args' - nparam) + (mkGHole ())))) + in + let eq' = + DAst.make ?loc:loc1 @@ GApp(DAst.make ?loc:loc2 @@GRef(jmeq,None),[ty;DAst.make ?loc:loc3 @@ GVar id;rt_typ;rt]) + in observe (str "computing new type for jmeq : " ++ pr_glob_constr_env env eq'); - let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in - observe (str " computing new type for jmeq : done") ; + let eq'_as_constr,ctx = Pretyping.understand env (Evd.from_env env) eq' in + observe (str " computing new type for jmeq : done") ; let sigma = Evd.(from_env env) in - let new_args = + let new_args = match EConstr.kind sigma eq'_as_constr with - | App(_,[|_;_;ty;_|]) -> + | App(_,[|_;_;ty;_|]) -> let ty = Array.to_list (snd (EConstr.destApp sigma ty)) in - let ty' = snd (Util.List.chop nparam ty) in - List.fold_left2 - (fun acc var_as_constr arg -> - if isRel var_as_constr - then - let na = RelDecl.get_name (Environ.lookup_rel (destRel var_as_constr) env) in - match na with - | Anonymous -> acc - | Name id' -> - (id',Detyping.detype Detyping.Now false Id.Set.empty - env + let ty' = snd (Util.List.chop nparam ty) in + List.fold_left2 + (fun acc var_as_constr arg -> + if isRel var_as_constr + then + let na = RelDecl.get_name (Environ.lookup_rel (destRel var_as_constr) env) in + match na with + | Anonymous -> acc + | Name id' -> + (id',Detyping.detype Detyping.Now false Id.Set.empty + env (Evd.from_env env) - arg)::acc - else if isVar var_as_constr - then (destVar var_as_constr,Detyping.detype Detyping.Now false Id.Set.empty - env + arg)::acc + else if isVar var_as_constr + then (destVar var_as_constr,Detyping.detype Detyping.Now false Id.Set.empty + env (Evd.from_env env) - arg)::acc - else acc - ) - [] - arg' - ty' - | _ -> assert false - in - let is_in_b = is_free_in id b in - let _keep_eq = - not (List.exists (is_free_in id) args) || is_in_b || - List.exists (is_free_in id) crossed_types - in - let new_args = - List.fold_left - (fun args (id,rt) -> - List.map (replace_var_by_term id rt) args - ) - args - ((id,rt)::new_args) - in - let subst_b = - if is_in_b then b else replace_var_by_term id rt b - in - let new_env = - let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in + arg)::acc + else acc + ) + [] + arg' + ty' + | _ -> assert false + in + let is_in_b = is_free_in id b in + let _keep_eq = + not (List.exists (is_free_in id) args) || is_in_b || + List.exists (is_free_in id) crossed_types + in + let new_args = + List.fold_left + (fun args (id,rt) -> + List.map (replace_var_by_term id rt) args + ) + args + ((id,rt)::new_args) + in + let subst_b = + if is_in_b then b else replace_var_by_term id rt b + in + let new_env = + let t',ctx = Pretyping.understand env (Evd.from_env env) eq' in let r = Sorts.Relevant in (* TODO relevance *) EConstr.push_rel (LocalAssum (make_annot n r,t')) env in - let new_b,id_to_exclude = - rebuild_cons - new_env - nb_args relname - new_args new_crossed_types - (depth + 1) subst_b - in - mkGProd(n,eq',new_b),id_to_exclude - end - (* J.F:. keep this comment it explain how to remove some meaningless equalities - if keep_eq then - mkGProd(n,t,new_b),id_to_exclude - else new_b, Id.Set.add id id_to_exclude - *) - | GApp(eq_as_ref,[ty;rt1;rt2]) + let new_b,id_to_exclude = + rebuild_cons + new_env + nb_args relname + new_args new_crossed_types + (depth + 1) subst_b + in + mkGProd(n,eq',new_b),id_to_exclude + end + (* J.F:. keep this comment it explain how to remove some meaningless equalities + if keep_eq then + mkGProd(n,t,new_b),id_to_exclude + else new_b, Id.Set.add id id_to_exclude + *) + | GApp(eq_as_ref,[ty;rt1;rt2]) when is_gr eq_as_ref Coqlib.(lib_ref "core.eq.type") && n == Anonymous - -> - begin - try + -> + begin + try let l = decompose_raw_eq env rt1 rt2 in - if List.length l > 1 - then - let new_rt = - List.fold_left - (fun acc (lhs,rhs) -> - mkGProd(Anonymous, + if List.length l > 1 + then + let new_rt = + List.fold_left + (fun acc (lhs,rhs) -> + mkGProd(Anonymous, mkGApp(mkGRef Coqlib.(lib_ref "core.eq.type"),[mkGHole ();lhs;rhs]),acc) - ) - b - l - in - rebuild_cons env nb_args relname args crossed_types depth new_rt - else raise Continue - with Continue -> + ) + b + l + in + rebuild_cons env nb_args relname args crossed_types depth new_rt + else raise Continue + with Continue -> observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt); - let t',ctx = Pretyping.understand env (Evd.from_env env) t in + let t',ctx = Pretyping.understand env (Evd.from_env env) t in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = - rebuild_cons new_env - nb_args relname - args new_crossed_types - (depth + 1) b - in - match n with - | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> - new_b,Id.Set.remove id - (Id.Set.filter not_free_in_t id_to_exclude) - | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude - end - | _ -> + rebuild_cons new_env + nb_args relname + args new_crossed_types + (depth + 1) b + in + match n with + | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> + new_b,Id.Set.remove id + (Id.Set.filter not_free_in_t id_to_exclude) + | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude + end + | _ -> observe (str "computing new type for prod : " ++ pr_glob_constr_env env rt); - let t',ctx = Pretyping.understand env (Evd.from_env env) t in + let t',ctx = Pretyping.understand env (Evd.from_env env) t in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = - rebuild_cons new_env - nb_args relname - args new_crossed_types - (depth + 1) b - in - match n with - | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> - new_b,Id.Set.remove id - (Id.Set.filter not_free_in_t id_to_exclude) - | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude - end + rebuild_cons new_env + nb_args relname + args new_crossed_types + (depth + 1) b + in + match n with + | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> + new_b,Id.Set.remove id + (Id.Set.filter not_free_in_t id_to_exclude) + | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude + end | GLambda(n,k,t,b) -> - begin - let not_free_in_t id = not (is_free_in id t) in - let new_crossed_types = t :: crossed_types in + begin + let not_free_in_t id = not (is_free_in id t) in + let new_crossed_types = t :: crossed_types in observe (str "computing new type for lambda : " ++ pr_glob_constr_env env rt); - let t',ctx = Pretyping.understand env (Evd.from_env env) t in - match n with - | Name id -> + let t',ctx = Pretyping.understand env (Evd.from_env env) t in + match n with + | Name id -> let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot n r,t')) env in let new_b,id_to_exclude = - rebuild_cons new_env - nb_args relname - (args@[mkGVar id])new_crossed_types - (depth + 1 ) b - in - if Id.Set.mem id id_to_exclude && depth >= nb_args - then - new_b, Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) - else - DAst.make @@ GProd(n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude - | _ -> anomaly (Pp.str "Should not have an anonymous function here.") - (* We have renamed all the anonymous functions during alpha_renaming phase *) - - end + rebuild_cons new_env + nb_args relname + (args@[mkGVar id])new_crossed_types + (depth + 1 ) b + in + if Id.Set.mem id id_to_exclude && depth >= nb_args + then + new_b, Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) + else + DAst.make @@ GProd(n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude + | _ -> anomaly (Pp.str "Should not have an anonymous function here.") + (* We have renamed all the anonymous functions during alpha_renaming phase *) + + end | GLetIn(n,v,t,b) -> - begin + begin let t = match t with None -> v | Some t -> DAst.make ?loc:rt.loc @@ GCast (v,CastConv t) in - let not_free_in_t id = not (is_free_in id t) in - let evd = (Evd.from_env env) in - let t',ctx = Pretyping.understand env evd t in - let evd = Evd.from_ctx ctx in + let not_free_in_t id = not (is_free_in id t) in + let evd = (Evd.from_env env) in + let t',ctx = Pretyping.understand env evd t in + let evd = Evd.from_ctx ctx in let type_t' = Typing.unsafe_type_of env evd t' in let t' = EConstr.Unsafe.to_constr t' in - let type_t' = EConstr.Unsafe.to_constr type_t' in + let type_t' = EConstr.Unsafe.to_constr type_t' in let new_env = Environ.push_rel (LocalDef (make_annot n Sorts.Relevant,t',type_t')) env in - let new_b,id_to_exclude = - rebuild_cons new_env - nb_args relname - args (t::crossed_types) - (depth + 1 ) b in - match n with - | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> - new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) - | _ -> DAst.make @@ GLetIn(n,t,None,new_b), (* HOPING IT WOULD WORK *) - Id.Set.filter not_free_in_t id_to_exclude - end + let new_b,id_to_exclude = + rebuild_cons new_env + nb_args relname + args (t::crossed_types) + (depth + 1 ) b in + match n with + | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args -> + new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) + | _ -> DAst.make @@ GLetIn(n,t,None,new_b), (* HOPING IT WOULD WORK *) + Id.Set.filter not_free_in_t id_to_exclude + end | GLetTuple(nal,(na,rto),t,b) -> - assert (Option.is_empty rto); - begin - let not_free_in_t id = not (is_free_in id t) in - let new_t,id_to_exclude' = - rebuild_cons env - nb_args - relname - args (crossed_types) - depth t - in - let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in + assert (Option.is_empty rto); + begin + let not_free_in_t id = not (is_free_in id t) in + let new_t,id_to_exclude' = + rebuild_cons env + nb_args + relname + args (crossed_types) + depth t + in + let t',ctx = Pretyping.understand env (Evd.from_env env) new_t in let r = Sorts.Relevant in (* TODO relevance *) let new_env = EConstr.push_rel (LocalAssum (make_annot na r,t')) env in let new_b,id_to_exclude = - rebuild_cons new_env - nb_args relname - args (t::crossed_types) - (depth + 1) b - in + rebuild_cons new_env + nb_args relname + args (t::crossed_types) + (depth + 1) b + in (* match n with *) (* | Name id when Id.Set.mem id id_to_exclude -> *) (* new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) *) (* | _ -> *) - DAst.make @@ GLetTuple(nal,(na,None),t,new_b), - Id.Set.filter not_free_in_t (Id.Set.union id_to_exclude id_to_exclude') + DAst.make @@ GLetTuple(nal,(na,None),t,new_b), + Id.Set.filter not_free_in_t (Id.Set.union id_to_exclude id_to_exclude') - end + end | _ -> mkGApp(mkGVar relname,args@[rt]),Id.Set.empty @@ -1249,7 +1249,7 @@ let rec compute_cst_params relnames params gt = DAst.with_val (function compute_cst_params relnames t_params b | GCases _ -> params (* If there is still cases at this point they can only be - discrimination ones *) + discrimination ones *) | GSort _ -> params | GHole _ -> params | GIf _ | GRec _ | GCast _ -> @@ -1260,17 +1260,17 @@ and compute_cst_params_from_app acc (params,rtl) = match params,rtl with | _::_,[] -> assert false (* the rel has at least nargs + 1 arguments ! *) | ((Name id,_,None) as param)::params', c::rtl' when is_gid id c -> - compute_cst_params_from_app (param::acc) (params',rtl') + compute_cst_params_from_app (param::acc) (params',rtl') | _ -> List.rev acc let compute_params_name relnames (args : (Name.t * Glob_term.glob_constr * glob_constr option) list array) csts = let rels_params = Array.mapi (fun i args -> - List.fold_left - (fun params (_,cst) -> compute_cst_params relnames params cst) - args - csts.(i) + List.fold_left + (fun params (_,cst) -> compute_cst_params relnames params cst) + args + csts.(i) ) args in @@ -1278,16 +1278,16 @@ let compute_params_name relnames (args : (Name.t * Glob_term.glob_constr * glob_ let _ = try List.iteri - (fun i ((n,nt,typ) as param) -> - if Array.for_all - (fun l -> - let (n',nt',typ') = List.nth l i in - Name.equal n n' && glob_constr_eq nt nt' && Option.equal glob_constr_eq typ typ') - rels_params - then - l := param::!l - ) - rels_params.(0) + (fun i ((n,nt,typ) as param) -> + if Array.for_all + (fun l -> + let (n',nt',typ') = List.nth l i in + Name.equal n n' && glob_constr_eq nt nt' && Option.equal glob_constr_eq typ typ') + rels_params + then + l := param::!l + ) + rels_params.(0) with e when CErrors.noncritical e -> () in @@ -1333,7 +1333,7 @@ let do_build_inductive let t = EConstr.Unsafe.to_constr t in evd, Environ.push_named (LocalAssum (make_annot id Sorts.Relevant,t)) - env + env ) funnames (Array.of_list funconstants) @@ -1350,23 +1350,23 @@ let do_build_inductive let env_with_graphs = let rel_arity i funargs = (* Rebuilding arities (with parameters) *) let rel_first_args :(Name.t * Glob_term.glob_constr * Glob_term.glob_constr option ) list = - funargs + funargs in List.fold_right - (fun (n,t,typ) acc -> + (fun (n,t,typ) acc -> match typ with | Some typ -> CAst.make @@ Constrexpr.CLetIn((CAst.make n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ), - acc) - | None -> - CAst.make @@ Constrexpr.CProdN + acc) + | None -> + CAst.make @@ Constrexpr.CProdN ([Constrexpr.CLocalAssum([CAst.make n],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t)], - acc - ) - ) - rel_first_args - (rebuild_return_type returned_types.(i)) + acc + ) + ) + rel_first_args + (rebuild_return_type returned_types.(i)) in (* We need to lift back our work topconstr but only with all information We mimic a Set Printing All. @@ -1383,15 +1383,15 @@ let do_build_inductive let constr i res = List.map (function result (* (args',concl') *) -> - let rt = compose_glob_context result.context result.value in - let nb_args = List.length funsargs.(i) in - (* with_full_print (fun rt -> Pp.msgnl (str "glob constr " ++ pr_glob_constr rt)) rt; *) - fst ( - rebuild_cons env_with_graphs nb_args relnames.(i) - [] - [] - rt - ) + let rt = compose_glob_context result.context result.value in + let nb_args = List.length funsargs.(i) in + (* with_full_print (fun rt -> Pp.msgnl (str "glob constr " ++ pr_glob_constr rt)) rt; *) + fst ( + rebuild_cons env_with_graphs nb_args relnames.(i) + [] + [] + rt + ) ) res.result in @@ -1427,12 +1427,12 @@ let do_build_inductive | Some typ -> CAst.make @@ Constrexpr.CLetIn((CAst.make n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ), - acc) - | None -> + acc) + | None -> CAst.make @@ Constrexpr.CProdN ([Constrexpr.CLocalAssum([CAst.make n],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t)], - acc - ) + acc + ) ) rel_first_args (rebuild_return_type returned_types.(i)) @@ -1446,7 +1446,7 @@ let do_build_inductive List.fold_left (fun acc (na,_,_) -> match na with - Anonymous -> acc + Anonymous -> acc | Name id -> id::acc ) [] @@ -1459,8 +1459,8 @@ let do_build_inductive | Some typ -> Constrexpr.CLocalDef((CAst.make n), Constrextern.extern_glob_constr Id.Set.empty t, Some (with_full_print (Constrextern.extern_glob_constr Id.Set.empty) typ)) - | None -> - Constrexpr.CLocalAssum + | None -> + Constrexpr.CLocalAssum ([(CAst.make n)], Constrexpr_ops.default_binder_kind, Constrextern.extern_glob_constr Id.Set.empty t) ) rels_params @@ -1469,9 +1469,9 @@ let do_build_inductive Array.map (List.map (fun (id,t) -> false,((CAst.make id), - with_full_print - (Constrextern.extern_glob_type Id.Set.empty) ((* zeta_normalize *) (alpha_rt rel_params_ids t)) - ) + with_full_print + (Constrextern.extern_glob_type Id.Set.empty) ((* zeta_normalize *) (alpha_rt rel_params_ids t)) + ) )) (rel_constructors) in @@ -1510,35 +1510,35 @@ let do_build_inductive Declarations.Finite with | UserError(s,msg) as e -> - let _time3 = System.get_time () in + let _time3 = System.get_time () in (* Pp.msgnl (str "error : "++ str (string_of_float (System.time_difference time2 time3))); *) - let repacked_rel_inds = + let repacked_rel_inds = List.map (fun ((a , b , c , l),ntn) -> ((false,(a,None)) , b, c , Vernacexpr.Inductive_kw, Vernacexpr.Constructors l),ntn ) - rel_inds - in - let msg = - str "while trying to define"++ spc () ++ + rel_inds + in + let msg = + str "while trying to define"++ spc () ++ Ppvernac.pr_vernac (CAst.make Vernacexpr.{ control = []; attrs = []; expr = VernacInductive(None,false,Declarations.Finite,repacked_rel_inds)}) - ++ fnl () ++ - msg - in - observe (msg); - raise e + ++ fnl () ++ + msg + in + observe (msg); + raise e | reraise -> - let _time3 = System.get_time () in + let _time3 = System.get_time () in (* Pp.msgnl (str "error : "++ str (string_of_float (System.time_difference time2 time3))); *) - let repacked_rel_inds = + let repacked_rel_inds = List.map (fun ((a , b , c , l),ntn) -> ((false,(a,None)) , b, c , Vernacexpr.Inductive_kw, Vernacexpr.Constructors l),ntn ) - rel_inds - in - let msg = - str "while trying to define"++ spc () ++ + rel_inds + in + let msg = + str "while trying to define"++ spc () ++ Ppvernac.pr_vernac (CAst.make @@ Vernacexpr.{ control = []; attrs = []; expr = VernacInductive(None,false,Declarations.Finite,repacked_rel_inds)}) - ++ fnl () ++ - CErrors.print reraise - in - observe msg; - raise reraise + ++ fnl () ++ + CErrors.print reraise + in + observe msg; + raise reraise diff --git a/plugins/funind/glob_term_to_relation.mli b/plugins/funind/glob_term_to_relation.mli index ff0e98d00f..a29e5dff23 100644 --- a/plugins/funind/glob_term_to_relation.mli +++ b/plugins/funind/glob_term_to_relation.mli @@ -8,10 +8,10 @@ open Names val build_inductive : (* (ModPath.t * DirPath.t) option -> - Id.t list -> (* The list of function name *) + Id.t list -> (* The list of function name *) *) Evd.evar_map -> - Constr.pconstant list -> + Constr.pconstant list -> (Name.t*Glob_term.glob_constr*Glob_term.glob_constr option) list list -> (* The list of function args *) Constrexpr.constr_expr list -> (* The list of function returned type *) Glob_term.glob_constr list -> (* the list of body *) diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml index 5f54bad598..f2d98a13ab 100644 --- a/plugins/funind/glob_termops.ml +++ b/plugins/funind/glob_termops.ml @@ -36,7 +36,7 @@ let glob_decompose_app = (* msgnl (str "glob_decompose_app on : "++ Printer.pr_glob_constr rt); *) match DAst.get rt with | GApp(rt,rtl) -> - decompose_rapp (List.fold_left (fun y x -> x::y) acc rtl) rt + decompose_rapp (List.fold_left (fun y x -> x::y) acc rtl) rt | _ -> rt,List.rev acc in decompose_rapp [] @@ -62,62 +62,62 @@ let change_vars = DAst.map_with_loc (fun ?loc -> function | GRef _ as x -> x | GVar id -> - let new_id = - try - Id.Map.find id mapping - with Not_found -> id - in - GVar(new_id) + let new_id = + try + Id.Map.find id mapping + with Not_found -> id + in + GVar(new_id) | GEvar _ as x -> x | GPatVar _ as x -> x | GApp(rt',rtl) -> - GApp(change_vars mapping rt', - List.map (change_vars mapping) rtl - ) + GApp(change_vars mapping rt', + List.map (change_vars mapping) rtl + ) | GLambda(name,k,t,b) -> - GLambda(name, - k, - change_vars mapping t, - change_vars (remove_name_from_mapping mapping name) b - ) + GLambda(name, + k, + change_vars mapping t, + change_vars (remove_name_from_mapping mapping name) b + ) | GProd(name,k,t,b) -> - GProd( name, - k, - change_vars mapping t, - change_vars (remove_name_from_mapping mapping name) b - ) + GProd( name, + k, + change_vars mapping t, + change_vars (remove_name_from_mapping mapping name) b + ) | GLetIn(name,def,typ,b) -> - GLetIn(name, - change_vars mapping def, - Option.map (change_vars mapping) typ, - change_vars (remove_name_from_mapping mapping name) b - ) + GLetIn(name, + change_vars mapping def, + Option.map (change_vars mapping) typ, + change_vars (remove_name_from_mapping mapping name) b + ) | GLetTuple(nal,(na,rto),b,e) -> - let new_mapping = List.fold_left remove_name_from_mapping mapping nal in - GLetTuple(nal, - (na, Option.map (change_vars mapping) rto), - change_vars mapping b, - change_vars new_mapping e - ) + let new_mapping = List.fold_left remove_name_from_mapping mapping nal in + GLetTuple(nal, + (na, Option.map (change_vars mapping) rto), + change_vars mapping b, + change_vars new_mapping e + ) | GCases(sty,infos,el,brl) -> - GCases(sty, - infos, - List.map (fun (e,x) -> (change_vars mapping e,x)) el, - List.map (change_vars_br mapping) brl - ) + GCases(sty, + infos, + List.map (fun (e,x) -> (change_vars mapping e,x)) el, + List.map (change_vars_br mapping) brl + ) | GIf(b,(na,e_option),lhs,rhs) -> - GIf(change_vars mapping b, - (na,Option.map (change_vars mapping) e_option), - change_vars mapping lhs, - change_vars mapping rhs - ) + GIf(change_vars mapping b, + (na,Option.map (change_vars mapping) e_option), + change_vars mapping lhs, + change_vars mapping rhs + ) | GRec _ -> user_err ?loc Pp.(str "Local (co)fixes are not supported") | GSort _ as x -> x | GHole _ as x -> x | GInt _ as x -> x | GFloat _ as x -> x | GCast(b,c) -> - GCast(change_vars mapping b, + GCast(change_vars mapping b, Glob_ops.map_cast_type (change_vars mapping) c) ) rt and change_vars_br mapping ({CAst.loc;v=(idl,patl,res)} as br) = @@ -134,40 +134,40 @@ let rec alpha_pat excluded pat = let loc = pat.CAst.loc in match DAst.get pat with | PatVar Anonymous -> - let new_id = Indfun_common.fresh_id excluded "_x" in - (DAst.make ?loc @@ PatVar(Name new_id)),(new_id::excluded),Id.Map.empty + let new_id = Indfun_common.fresh_id excluded "_x" in + (DAst.make ?loc @@ PatVar(Name new_id)),(new_id::excluded),Id.Map.empty | PatVar(Name id) -> - if Id.List.mem id excluded - then - let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in - (DAst.make ?loc @@ PatVar(Name new_id)),(new_id::excluded), - (Id.Map.add id new_id Id.Map.empty) - else pat, excluded,Id.Map.empty + if Id.List.mem id excluded + then + let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in + (DAst.make ?loc @@ PatVar(Name new_id)),(new_id::excluded), + (Id.Map.add id new_id Id.Map.empty) + else pat, excluded,Id.Map.empty | PatCstr(constr,patl,na) -> - let new_na,new_excluded,map = - match na with - | Name id when Id.List.mem id excluded -> - let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in - Name new_id,new_id::excluded, Id.Map.add id new_id Id.Map.empty - | _ -> na,excluded,Id.Map.empty - in - let new_patl,new_excluded,new_map = - List.fold_left - (fun (patl,excluded,map) pat -> - let new_pat,new_excluded,new_map = alpha_pat excluded pat in - (new_pat::patl,new_excluded,Id.Map.fold Id.Map.add new_map map) - ) - ([],new_excluded,map) - patl - in + let new_na,new_excluded,map = + match na with + | Name id when Id.List.mem id excluded -> + let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in + Name new_id,new_id::excluded, Id.Map.add id new_id Id.Map.empty + | _ -> na,excluded,Id.Map.empty + in + let new_patl,new_excluded,new_map = + List.fold_left + (fun (patl,excluded,map) pat -> + let new_pat,new_excluded,new_map = alpha_pat excluded pat in + (new_pat::patl,new_excluded,Id.Map.fold Id.Map.add new_map map) + ) + ([],new_excluded,map) + patl + in (DAst.make ?loc @@ PatCstr(constr,List.rev new_patl,new_na)),new_excluded,new_map let alpha_patl excluded patl = let patl,new_excluded,map = List.fold_left (fun (patl,excluded,map) pat -> - let new_pat,new_excluded,new_map = alpha_pat excluded pat in - new_pat::patl,new_excluded,(Id.Map.fold Id.Map.add new_map map) + let new_pat,new_excluded,new_map = alpha_pat excluded pat in + new_pat::patl,new_excluded,(Id.Map.fold Id.Map.add new_map map) ) ([],excluded,Id.Map.empty) patl @@ -182,15 +182,15 @@ let raw_get_pattern_id pat acc = match DAst.get pat with | PatVar(Anonymous) -> assert false | PatVar(Name id) -> - [id] + [id] | PatCstr(constr,patternl,_) -> - List.fold_right - (fun pat idl -> - let idl' = get_pattern_id pat in - idl'@idl - ) - patternl - [] + List.fold_right + (fun pat idl -> + let idl' = get_pattern_id pat in + idl'@idl + ) + patternl + [] in (get_pattern_id pat)@acc @@ -202,109 +202,109 @@ let rec alpha_rt excluded rt = match DAst.get rt with | GRef _ | GVar _ | GEvar _ | GPatVar _ as rt -> rt | GLambda(Anonymous,k,t,b) -> - let new_id = Namegen.next_ident_away (Id.of_string "_x") (Id.Set.of_list excluded) in - let new_excluded = new_id :: excluded in - let new_t = alpha_rt new_excluded t in - let new_b = alpha_rt new_excluded b in - GLambda(Name new_id,k,new_t,new_b) + let new_id = Namegen.next_ident_away (Id.of_string "_x") (Id.Set.of_list excluded) in + let new_excluded = new_id :: excluded in + let new_t = alpha_rt new_excluded t in + let new_b = alpha_rt new_excluded b in + GLambda(Name new_id,k,new_t,new_b) | GProd(Anonymous,k,t,b) -> - let new_t = alpha_rt excluded t in - let new_b = alpha_rt excluded b in - GProd(Anonymous,k,new_t,new_b) + let new_t = alpha_rt excluded t in + let new_b = alpha_rt excluded b in + GProd(Anonymous,k,new_t,new_b) | GLetIn(Anonymous,b,t,c) -> - let new_b = alpha_rt excluded b in - let new_t = Option.map (alpha_rt excluded) t in - let new_c = alpha_rt excluded c in - GLetIn(Anonymous,new_b,new_t,new_c) + let new_b = alpha_rt excluded b in + let new_t = Option.map (alpha_rt excluded) t in + let new_c = alpha_rt excluded c in + GLetIn(Anonymous,new_b,new_t,new_c) | GLambda(Name id,k,t,b) -> - let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in - let t,b = - if Id.equal new_id id - then t, b - else - let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in - (t,replace b) - in - let new_excluded = new_id::excluded in - let new_t = alpha_rt new_excluded t in - let new_b = alpha_rt new_excluded b in - GLambda(Name new_id,k,new_t,new_b) + let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in + let t,b = + if Id.equal new_id id + then t, b + else + let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in + (t,replace b) + in + let new_excluded = new_id::excluded in + let new_t = alpha_rt new_excluded t in + let new_b = alpha_rt new_excluded b in + GLambda(Name new_id,k,new_t,new_b) | GProd(Name id,k,t,b) -> - let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in - let new_excluded = new_id::excluded in - let t,b = - if Id.equal new_id id - then t,b - else - let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in - (t,replace b) - in - let new_t = alpha_rt new_excluded t in - let new_b = alpha_rt new_excluded b in - GProd(Name new_id,k,new_t,new_b) + let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in + let new_excluded = new_id::excluded in + let t,b = + if Id.equal new_id id + then t,b + else + let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in + (t,replace b) + in + let new_t = alpha_rt new_excluded t in + let new_b = alpha_rt new_excluded b in + GProd(Name new_id,k,new_t,new_b) | GLetIn(Name id,b,t,c) -> - let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in - let c = - if Id.equal new_id id then c - else change_vars (Id.Map.add id new_id Id.Map.empty) c - in - let new_excluded = new_id::excluded in - let new_b = alpha_rt new_excluded b in - let new_t = Option.map (alpha_rt new_excluded) t in - let new_c = alpha_rt new_excluded c in - GLetIn(Name new_id,new_b,new_t,new_c) + let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in + let c = + if Id.equal new_id id then c + else change_vars (Id.Map.add id new_id Id.Map.empty) c + in + let new_excluded = new_id::excluded in + let new_b = alpha_rt new_excluded b in + let new_t = Option.map (alpha_rt new_excluded) t in + let new_c = alpha_rt new_excluded c in + GLetIn(Name new_id,new_b,new_t,new_c) | GLetTuple(nal,(na,rto),t,b) -> - let rev_new_nal,new_excluded,mapping = - List.fold_left - (fun (nal,excluded,mapping) na -> - match na with - | Anonymous -> (na::nal,excluded,mapping) - | Name id -> - let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in - if Id.equal new_id id - then - na::nal,id::excluded,mapping - else - (Name new_id)::nal,id::excluded,(Id.Map.add id new_id mapping) - ) - ([],excluded,Id.Map.empty) - nal - in - let new_nal = List.rev rev_new_nal in - let new_rto,new_t,new_b = - if Id.Map.is_empty mapping - then rto,t,b - else let replace = change_vars mapping in - (Option.map replace rto, t,replace b) - in - let new_t = alpha_rt new_excluded new_t in - let new_b = alpha_rt new_excluded new_b in - let new_rto = Option.map (alpha_rt new_excluded) new_rto in - GLetTuple(new_nal,(na,new_rto),new_t,new_b) + let rev_new_nal,new_excluded,mapping = + List.fold_left + (fun (nal,excluded,mapping) na -> + match na with + | Anonymous -> (na::nal,excluded,mapping) + | Name id -> + let new_id = Namegen.next_ident_away id (Id.Set.of_list excluded) in + if Id.equal new_id id + then + na::nal,id::excluded,mapping + else + (Name new_id)::nal,id::excluded,(Id.Map.add id new_id mapping) + ) + ([],excluded,Id.Map.empty) + nal + in + let new_nal = List.rev rev_new_nal in + let new_rto,new_t,new_b = + if Id.Map.is_empty mapping + then rto,t,b + else let replace = change_vars mapping in + (Option.map replace rto, t,replace b) + in + let new_t = alpha_rt new_excluded new_t in + let new_b = alpha_rt new_excluded new_b in + let new_rto = Option.map (alpha_rt new_excluded) new_rto in + GLetTuple(new_nal,(na,new_rto),new_t,new_b) | GCases(sty,infos,el,brl) -> - let new_el = - List.map (function (rt,i) -> alpha_rt excluded rt, i) el - in - GCases(sty,infos,new_el,List.map (alpha_br excluded) brl) + let new_el = + List.map (function (rt,i) -> alpha_rt excluded rt, i) el + in + GCases(sty,infos,new_el,List.map (alpha_br excluded) brl) | GIf(b,(na,e_o),lhs,rhs) -> - GIf(alpha_rt excluded b, - (na,Option.map (alpha_rt excluded) e_o), - alpha_rt excluded lhs, - alpha_rt excluded rhs - ) + GIf(alpha_rt excluded b, + (na,Option.map (alpha_rt excluded) e_o), + alpha_rt excluded lhs, + alpha_rt excluded rhs + ) | GRec _ -> user_err Pp.(str "Not handled GRec") | GSort _ | GInt _ | GFloat _ | GHole _ as rt -> rt | GCast (b,c) -> - GCast(alpha_rt excluded b, + GCast(alpha_rt excluded b, Glob_ops.map_cast_type (alpha_rt excluded) c) | GApp(f,args) -> - GApp(alpha_rt excluded f, - List.map (alpha_rt excluded) args - ) + GApp(alpha_rt excluded f, + List.map (alpha_rt excluded) args + ) in new_rt @@ -327,30 +327,30 @@ let is_free_in id = | GPatVar _ -> false | GApp(rt,rtl) -> List.exists is_free_in (rt::rtl) | GLambda(n,_,t,b) | GProd(n,_,t,b) -> - let check_in_b = - match n with - | Name id' -> not (Id.equal id' id) - | _ -> true - in - is_free_in t || (check_in_b && is_free_in b) + let check_in_b = + match n with + | Name id' -> not (Id.equal id' id) + | _ -> true + in + is_free_in t || (check_in_b && is_free_in b) | GLetIn(n,b,t,c) -> - let check_in_c = - match n with - | Name id' -> not (Id.equal id' id) - | _ -> true - in - is_free_in b || Option.cata is_free_in true t || (check_in_c && is_free_in c) + let check_in_c = + match n with + | Name id' -> not (Id.equal id' id) + | _ -> true + in + is_free_in b || Option.cata is_free_in true t || (check_in_c && is_free_in c) | GCases(_,_,el,brl) -> - (List.exists (fun (e,_) -> is_free_in e) el) || - List.exists is_free_in_br brl + (List.exists (fun (e,_) -> is_free_in e) el) || + List.exists is_free_in_br brl | GLetTuple(nal,_,b,t) -> - let check_in_nal = - not (List.exists (function Name id' -> Id.equal id' id | _ -> false) nal) - in - is_free_in t || (check_in_nal && is_free_in b) + let check_in_nal = + not (List.exists (function Name id' -> Id.equal id' id | _ -> false) nal) + in + is_free_in t || (check_in_nal && is_free_in b) | GIf(cond,_,br1,br2) -> - is_free_in cond || is_free_in br1 || is_free_in br2 + is_free_in cond || is_free_in br1 || is_free_in br2 | GRec _ -> user_err Pp.(str "Not handled GRec") | GSort _ -> false | GHole _ -> false @@ -368,26 +368,26 @@ let is_free_in id = let rec pattern_to_term pt = DAst.with_val (function | PatVar Anonymous -> assert false | PatVar(Name id) -> - mkGVar id + mkGVar id | PatCstr(constr,patternl,_) -> let cst_narg = Inductiveops.constructor_nallargs - (Global.env ()) - constr + (Global.env ()) + constr in let implicit_args = - Array.to_list - (Array.init - (cst_narg - List.length patternl) - (fun _ -> mkGHole ()) - ) + Array.to_list + (Array.init + (cst_narg - List.length patternl) + (fun _ -> mkGHole ()) + ) in let patl_as_term = - List.map pattern_to_term patternl + List.map pattern_to_term patternl in mkGApp(mkGRef(GlobRef.ConstructRef constr), - implicit_args@patl_as_term - ) + implicit_args@patl_as_term + ) ) pt @@ -399,51 +399,51 @@ let replace_var_by_term x_id term = | GEvar _ | GPatVar _ as rt -> rt | GApp(rt',rtl) -> - GApp(replace_var_by_pattern rt', - List.map replace_var_by_pattern rtl - ) + GApp(replace_var_by_pattern rt', + List.map replace_var_by_pattern rtl + ) | GLambda(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GLambda(name,k,t,b) -> - GLambda(name, - k, - replace_var_by_pattern t, - replace_var_by_pattern b - ) + GLambda(name, + k, + replace_var_by_pattern t, + replace_var_by_pattern b + ) | GProd(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GProd(name,k,t,b) -> - GProd( name, - k, - replace_var_by_pattern t, - replace_var_by_pattern b - ) + GProd( name, + k, + replace_var_by_pattern t, + replace_var_by_pattern b + ) | GLetIn(Name id,_,_,_) as rt when Id.compare id x_id == 0 -> rt | GLetIn(name,def,typ,b) -> - GLetIn(name, - replace_var_by_pattern def, - Option.map (replace_var_by_pattern) typ, - replace_var_by_pattern b - ) + GLetIn(name, + replace_var_by_pattern def, + Option.map (replace_var_by_pattern) typ, + replace_var_by_pattern b + ) | GLetTuple(nal,_,_,_) as rt - when List.exists (function Name id -> Id.equal id x_id | _ -> false) nal -> - rt + when List.exists (function Name id -> Id.equal id x_id | _ -> false) nal -> + rt | GLetTuple(nal,(na,rto),def,b) -> - GLetTuple(nal, - (na,Option.map replace_var_by_pattern rto), - replace_var_by_pattern def, - replace_var_by_pattern b - ) + GLetTuple(nal, + (na,Option.map replace_var_by_pattern rto), + replace_var_by_pattern def, + replace_var_by_pattern b + ) | GCases(sty,infos,el,brl) -> - GCases(sty, - infos, - List.map (fun (e,x) -> (replace_var_by_pattern e,x)) el, - List.map replace_var_by_pattern_br brl - ) + GCases(sty, + infos, + List.map (fun (e,x) -> (replace_var_by_pattern e,x)) el, + List.map replace_var_by_pattern_br brl + ) | GIf(b,(na,e_option),lhs,rhs) -> - GIf(replace_var_by_pattern b, - (na,Option.map replace_var_by_pattern e_option), - replace_var_by_pattern lhs, - replace_var_by_pattern rhs - ) + GIf(replace_var_by_pattern b, + (na,Option.map replace_var_by_pattern e_option), + replace_var_by_pattern lhs, + replace_var_by_pattern rhs + ) | GRec _ -> CErrors.user_err (Pp.str "Not handled GRec") | GSort _ @@ -451,7 +451,7 @@ let replace_var_by_term x_id term = | GInt _ as rt -> rt | GFloat _ as rt -> rt | GCast(b,c) -> - GCast(replace_var_by_pattern b, + GCast(replace_var_by_pattern b, Glob_ops.map_cast_type replace_var_by_pattern c) ) x and replace_var_by_pattern_br ({CAst.loc;v=(idl,patl,res)} as br) = @@ -471,16 +471,16 @@ let rec are_unifiable_aux = function | [] -> () | (l, r) ::eqs -> match DAst.get l, DAst.get r with - | PatVar _ ,_ | _, PatVar _-> are_unifiable_aux eqs - | PatCstr(constructor1,cpl1,_), PatCstr(constructor2,cpl2,_) -> - if not (eq_constructor constructor2 constructor1) - then raise NotUnifiable - else - let eqs' = - try (List.combine cpl1 cpl2) @ eqs - with Invalid_argument _ -> anomaly (Pp.str "are_unifiable_aux.") - in - are_unifiable_aux eqs' + | PatVar _ ,_ | _, PatVar _-> are_unifiable_aux eqs + | PatCstr(constructor1,cpl1,_), PatCstr(constructor2,cpl2,_) -> + if not (eq_constructor constructor2 constructor1) + then raise NotUnifiable + else + let eqs' = + try (List.combine cpl1 cpl2) @ eqs + with Invalid_argument _ -> anomaly (Pp.str "are_unifiable_aux.") + in + are_unifiable_aux eqs' let are_unifiable pat1 pat2 = try @@ -493,17 +493,17 @@ let rec eq_cases_pattern_aux = function | [] -> () | (l, r) ::eqs -> match DAst.get l, DAst.get r with - | PatVar _, PatVar _ -> eq_cases_pattern_aux eqs - | PatCstr(constructor1,cpl1,_), PatCstr(constructor2,cpl2,_) -> - if not (eq_constructor constructor2 constructor1) - then raise NotUnifiable - else - let eqs' = - try (List.combine cpl1 cpl2) @ eqs - with Invalid_argument _ -> anomaly (Pp.str "eq_cases_pattern_aux.") - in - eq_cases_pattern_aux eqs' - | _ -> raise NotUnifiable + | PatVar _, PatVar _ -> eq_cases_pattern_aux eqs + | PatCstr(constructor1,cpl1,_), PatCstr(constructor2,cpl2,_) -> + if not (eq_constructor constructor2 constructor1) + then raise NotUnifiable + else + let eqs' = + try (List.combine cpl1 cpl2) @ eqs + with Invalid_argument _ -> anomaly (Pp.str "eq_cases_pattern_aux.") + in + eq_cases_pattern_aux eqs' + | _ -> raise NotUnifiable let eq_cases_pattern pat1 pat2 = try @@ -528,50 +528,50 @@ let expand_as = match DAst.get rt with | PatVar _ -> map | PatCstr(_,patl,Name id) -> - Id.Map.add id (pattern_to_term rt) (List.fold_left add_as map patl) + Id.Map.add id (pattern_to_term rt) (List.fold_left add_as map patl) | PatCstr(_,patl,_) -> List.fold_left add_as map patl in let rec expand_as map = DAst.map (function | GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ | GInt _ | GFloat _ as rt -> rt | GVar id as rt -> - begin - try - DAst.get (Id.Map.find id map) - with Not_found -> rt - end + begin + try + DAst.get (Id.Map.find id map) + with Not_found -> rt + end | GApp(f,args) -> GApp(expand_as map f,List.map (expand_as map) args) | GLambda(na,k,t,b) -> GLambda(na,k,expand_as map t, expand_as map b) | GProd(na,k,t,b) -> GProd(na,k,expand_as map t, expand_as map b) | GLetIn(na,v,typ,b) -> GLetIn(na, expand_as map v,Option.map (expand_as map) typ,expand_as map b) | GLetTuple(nal,(na,po),v,b) -> - GLetTuple(nal,(na,Option.map (expand_as map) po), - expand_as map v, expand_as map b) + GLetTuple(nal,(na,Option.map (expand_as map) po), + expand_as map v, expand_as map b) | GIf(e,(na,po),br1,br2) -> - GIf(expand_as map e,(na,Option.map (expand_as map) po), - expand_as map br1, expand_as map br2) + GIf(expand_as map e,(na,Option.map (expand_as map) po), + expand_as map br1, expand_as map br2) | GRec _ -> user_err Pp.(str "Not handled GRec") | GCast(b,c) -> - GCast(expand_as map b, + GCast(expand_as map b, Glob_ops.map_cast_type (expand_as map) c) | GCases(sty,po,el,brl) -> - GCases(sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el, - List.map (expand_as_br map) brl) + GCases(sty, Option.map (expand_as map) po, List.map (fun (rt,t) -> expand_as map rt,t) el, + List.map (expand_as_br map) brl) ) and expand_as_br map {CAst.loc; v=(idl,cpl,rt)} = CAst.make ?loc (idl,cpl, expand_as (List.fold_left add_as map cpl) rt) in expand_as Id.Map.empty -(* [resolve_and_replace_implicits ?expected_type env sigma rt] solves implicits of [rt] w.r.t. [env] and [sigma] and then replace them by their solution +(* [resolve_and_replace_implicits ?expected_type env sigma rt] solves implicits of [rt] w.r.t. [env] and [sigma] and then replace them by their solution *) exception Found of Evd.evar_info let resolve_and_replace_implicits ?(flags=Pretyping.all_and_fail_flags) ?(expected_type=Pretyping.WithoutTypeConstraint) env sigma rt = let open Evd in - let open Evar_kinds in + let open Evar_kinds in (* we first (pseudo) understand [rt] and get back the computed evar_map *) - (* FIXME : JF (30/03/2017) I'm not completely sure to have split understand as needed. -If someone knows how to prevent solved existantial removal in understand, please do not hesitate to change the computation of [ctx] here *) + (* FIXME : JF (30/03/2017) I'm not completely sure to have split understand as needed. +If someone knows how to prevent solved existantial removal in understand, please do not hesitate to change the computation of [ctx] here *) let ctx,_,_ = Pretyping.ise_pretype_gen flags env sigma Glob_ops.empty_lvar expected_type rt in let ctx = Evd.minimize_universes ctx in let f c = EConstr.of_constr (Evarutil.nf_evars_universes ctx (EConstr.Unsafe.to_constr c)) in @@ -603,7 +603,7 @@ If someone knows how to prevent solved existantial removal in understand, pleas ) | (GHole(BinderType na,_,_)) -> (* we only want to deal with implicit arguments *) ( - let res = + let res = try (* we scan the new evar map to find the evar corresponding to this hole (by looking the source *) Evd.fold (* to simulate an iter *) (fun _ evi _ -> @@ -622,9 +622,9 @@ If someone knows how to prevent solved existantial removal in understand, pleas (* we just have to lift the solution in glob_term *) Detyping.detype Detyping.Now false Id.Set.empty env ctx (f c) | Evar_empty -> rt (* the hole was not solved : we d when falseo nothing *) - in + in res ) - | _ -> Glob_ops.map_glob_constr change rt + | _ -> Glob_ops.map_glob_constr change rt in change rt diff --git a/plugins/funind/glob_termops.mli b/plugins/funind/glob_termops.mli index 70211a1860..bdde66bbd7 100644 --- a/plugins/funind/glob_termops.mli +++ b/plugins/funind/glob_termops.mli @@ -101,8 +101,8 @@ val ids_of_pat : cases_pattern -> Id.Set.t val expand_as : glob_constr -> glob_constr -(* [resolve_and_replace_implicits ?expected_type env sigma rt] solves implicits of [rt] w.r.t. [env] and [sigma] and then replace them by their solution +(* [resolve_and_replace_implicits ?expected_type env sigma rt] solves implicits of [rt] w.r.t. [env] and [sigma] and then replace them by their solution *) val resolve_and_replace_implicits : - ?flags:Pretyping.inference_flags -> + ?flags:Pretyping.inference_flags -> ?expected_type:Pretyping.typing_constraint -> Environ.env -> Evd.evar_map -> glob_constr -> glob_constr diff --git a/plugins/ltac/evar_tactics.ml b/plugins/ltac/evar_tactics.ml index 5211bedd46..c87eb7c3c9 100644 --- a/plugins/ltac/evar_tactics.ml +++ b/plugins/ltac/evar_tactics.ml @@ -51,20 +51,20 @@ let instantiate_tac n c ido = let sigma = gl.sigma in let evl = match ido with - ConclLocation () -> evar_list sigma (pf_concl gl) + ConclLocation () -> evar_list sigma (pf_concl gl) | HypLocation (id,hloc) -> let decl = Environ.lookup_named id (pf_env gl) in - match hloc with - InHyp -> - (match decl with + match hloc with + InHyp -> + (match decl with | LocalAssum (_,typ) -> evar_list sigma (EConstr.of_constr typ) - | _ -> user_err Pp.(str "Please be more specific: in type or value?")) - | InHypTypeOnly -> - evar_list sigma (EConstr.of_constr (NamedDecl.get_type decl)) - | InHypValueOnly -> - (match decl with + | _ -> user_err Pp.(str "Please be more specific: in type or value?")) + | InHypTypeOnly -> + evar_list sigma (EConstr.of_constr (NamedDecl.get_type decl)) + | InHypValueOnly -> + (match decl with | LocalDef (_,body,_) -> evar_list sigma (EConstr.of_constr body) - | _ -> user_err Pp.(str "Not a defined hypothesis.")) in + | _ -> user_err Pp.(str "Not a defined hypothesis.")) in if List.length evl < n then user_err Pp.(str "Not enough uninstantiated existential variables."); if n <= 0 then user_err Pp.(str "Incorrect existential variable index."); @@ -97,7 +97,7 @@ let let_evar name typ = Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma) (Tactics.pose_tac (Name.Name id) evar) end - + let hget_evar n = let open EConstr in Proofview.Goal.enter begin fun gl -> diff --git a/plugins/ltac/extraargs.mlg b/plugins/ltac/extraargs.mlg index e6e6e29d4f..bab6bfd78e 100644 --- a/plugins/ltac/extraargs.mlg +++ b/plugins/ltac/extraargs.mlg @@ -115,8 +115,8 @@ let interp_occs ist gl l = match l with | ArgArg x -> x | ArgVar ({ CAst.v = id } as locid) -> - (try int_list_of_VList (Id.Map.find id ist.lfun) - with Not_found | CannotCoerceTo _ -> [interp_int ist locid]) + (try int_list_of_VList (Id.Map.find id ist.lfun) + with Not_found | CannotCoerceTo _ -> [interp_int ist locid]) let interp_occs ist gl l = Tacmach.project gl , interp_occs ist gl l diff --git a/plugins/ltac/extratactics.mlg b/plugins/ltac/extratactics.mlg index f7215a9d13..a9e5271e81 100644 --- a/plugins/ltac/extratactics.mlg +++ b/plugins/ltac/extratactics.mlg @@ -204,7 +204,7 @@ TACTIC EXTEND dependent_rewrite END (** To be deprecated?, "cutrewrite (t=u) as <-" is equivalent to - "replace u with t" or "enough (t=u) as <-" and + "replace u with t" or "enough (t=u) as <-" and "cutrewrite (t=u) as ->" is equivalent to "enough (t=u) as ->". *) TACTIC EXTEND cut_rewrite @@ -314,8 +314,8 @@ let add_rewrite_hint ~poly bases ort t lcsr = let ctx = let ctx = UState.context_set ctx in if poly then ctx - else (* This is a global universe context that shouldn't be - refreshed at every use of the hint, declare it globally. *) + else (* This is a global universe context that shouldn't be + refreshed at every use of the hint, declare it globally. *) (Declare.declare_universe_context ~poly:false ctx; Univ.ContextSet.empty) in @@ -595,7 +595,7 @@ TACTIC EXTEND dep_generalize_eqs_vars | ["dependent" "generalize_eqs_vars" hyp(id) ] -> { abstract_generalize ~force_dep:true ~generalize_vars:true id } END -(** Tactic to automatically simplify hypotheses of the form [Π Δ, x_i = t_i -> T] +(** Tactic to automatically simplify hypotheses of the form [Π Δ, x_i = t_i -> T] where [t_i] is closed w.r.t. Δ. Such hypotheses are automatically generated during dependent induction. For internal use. *) @@ -613,17 +613,17 @@ END { -let subst_var_with_hole occ tid t = +let subst_var_with_hole occ tid t = let occref = if occ > 0 then ref occ else Find_subterm.error_invalid_occurrence [occ] in let locref = ref 0 in let rec substrec x = match DAst.get x with | GVar id -> - if Id.equal id tid + if Id.equal id tid then - (decr occref; - if Int.equal !occref 0 then x + (decr occref; + if Int.equal !occref 0 then x else - (incr locref; + (incr locref; DAst.make ~loc:(Loc.make_loc (!locref,0)) @@ GHole (Evar_kinds.QuestionMark { Evar_kinds.qm_obligation=Evar_kinds.Define true; @@ -648,7 +648,7 @@ let subst_hole_with_term occ tc t = decr occref; if Int.equal !occref 0 then tc else - (incr locref; + (incr locref; DAst.make ~loc:(Loc.make_loc (!locref,0)) @@ GHole (Evar_kinds.QuestionMark { Evar_kinds.qm_obligation=Evar_kinds.Define true; @@ -670,7 +670,7 @@ let hResolve id c occ t = let c_raw = Detyping.detype Detyping.Now true env_ids env sigma c in let t_raw = Detyping.detype Detyping.Now true env_ids env sigma t in let rec resolve_hole t_hole = - try + try Pretyping.understand env sigma t_hole with | Pretype_errors.PretypeError (_,_,Pretype_errors.UnsolvableImplicit _) as e -> @@ -686,7 +686,7 @@ let hResolve id c occ t = end let hResolve_auto id c t = - let rec resolve_auto n = + let rec resolve_auto n = try hResolve id c n t with @@ -727,7 +727,7 @@ exception Found of unit Proofview.tactic let rewrite_except h = Proofview.Goal.enter begin fun gl -> let hyps = Tacmach.New.pf_ids_of_hyps gl in - Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else + Tacticals.New.tclMAP (fun id -> if Id.equal id h then Proofview.tclUNIT () else Tacticals.New.tclTRY (Equality.general_rewrite_in true Locus.AllOccurrences true true id (mkVar h) false)) hyps end @@ -750,9 +750,9 @@ let mkCaseEq a : unit Proofview.tactic = (* FIXME: this looks really wrong. Does anybody really use this tactic? *) let (_, c) = Tacred.pattern_occs [Locus.OnlyOccurrences [1], a] env (Evd.from_env env) concl in - change_concl c + change_concl c end; - simplest_case a] + simplest_case a] end @@ -769,8 +769,8 @@ let case_eq_intros_rewrite x = let h = fresh_id_in_env hyps (Id.of_string "heq") (Proofview.Goal.env gl) in Tacticals.New.tclTHENLIST [ Tacticals.New.tclDO (n'-n-1) intro; - introduction h; - rewrite_except h] + introduction h; + rewrite_except h] end ] end @@ -781,14 +781,14 @@ let rec find_a_destructable_match sigma t = let dest = TacAtom (CAst.make @@ TacInductionDestruct(false, false, cl)) in match EConstr.kind sigma t with | Case (_,_,x,_) when closed0 sigma x -> - if isVar sigma x then - (* TODO check there is no rel n. *) - raise (Found (Tacinterp.eval_tactic dest)) - else - (* let _ = Pp.msgnl (Printer.pr_lconstr x) in *) - raise (Found (case_eq_intros_rewrite x)) + if isVar sigma x then + (* TODO check there is no rel n. *) + raise (Found (Tacinterp.eval_tactic dest)) + else + (* let _ = Pp.msgnl (Printer.pr_lconstr x) in *) + raise (Found (case_eq_intros_rewrite x)) | _ -> EConstr.iter sigma (fun c -> find_a_destructable_match sigma c) t - + let destauto t = Proofview.tclEVARMAP >>= fun sigma -> @@ -796,7 +796,7 @@ let destauto t = Tacticals.New.tclZEROMSG (str "No destructable match found") with Found tac -> tac -let destauto_in id = +let destauto_in id = Proofview.Goal.enter begin fun gl -> let ctype = Tacmach.New.pf_unsafe_type_of gl (mkVar id) in (* Pp.msgnl (Printer.pr_lconstr (mkVar id)); *) diff --git a/plugins/ltac/g_ltac.mlg b/plugins/ltac/g_ltac.mlg index cab8ed0a55..81a6651745 100644 --- a/plugins/ltac/g_ltac.mlg +++ b/plugins/ltac/g_ltac.mlg @@ -69,8 +69,8 @@ let test_bracket_ident = | KEYWORD "[" -> (match stream_nth 1 strm with | IDENT _ -> () - | _ -> raise Stream.Failure) - | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) + | _ -> raise Stream.Failure) (* Tactics grammar rules *) @@ -110,11 +110,11 @@ GRAMMAR EXTEND Gram | ta0 = tactic_expr; ";"; ta1 = tactic_expr -> { TacThen (ta0,ta1) } | ta0 = tactic_expr; ";"; l = tactic_then_locality; tg = tactic_then_gen; "]" -> { let (first,tail) = tg in - match l , tail with + match l , tail with | false , Some (t,last) -> TacThen (ta0,TacExtendTac (Array.of_list first, t, last)) - | true , Some (t,last) -> TacThens3parts (ta0, Array.of_list first, t, last) + | true , Some (t,last) -> TacThens3parts (ta0, Array.of_list first, t, last) | false , None -> TacThen (ta0,TacDispatch first) - | true , None -> TacThens (ta0,first) } ] + | true , None -> TacThens (ta0,first) } ] | "3" RIGHTA [ IDENT "try"; ta = tactic_expr -> { TacTry ta } | IDENT "do"; n = int_or_var; ta = tactic_expr -> { TacDo (n,ta) } @@ -148,12 +148,12 @@ GRAMMAR EXTEND Gram | b = match_key; c = tactic_expr; "with"; mrl = match_list; "end" -> { TacMatch (b,c,mrl) } | IDENT "first" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - { TacFirst l } + { TacFirst l } | IDENT "solve" ; "["; l = LIST0 tactic_expr SEP "|"; "]" -> - { TacSolve l } + { TacSolve l } | IDENT "idtac"; l = LIST0 message_token -> { TacId l } | g=failkw; n = [ n = int_or_var -> { n } | -> { fail_default_value } ]; - l = LIST0 message_token -> { TacFail (g,n,l) } + l = LIST0 message_token -> { TacFail (g,n,l) } | st = simple_tactic -> { st } | a = tactic_arg -> { TacArg(CAst.make ~loc a) } | r = reference; la = LIST0 tactic_arg_compat -> @@ -247,12 +247,12 @@ GRAMMAR EXTEND Gram | na = name; ":="; "["; mpv = match_pattern; "]"; ":"; mpt = match_pattern -> { Def (na, mpv, mpt) } | na = name; ":="; mpv = match_pattern -> { let t, ty = - match mpv with - | Term t -> (match t with - | { CAst.v = CCast (t, (CastConv ty | CastVM ty | CastNative ty)) } -> Term t, Some (Term ty) - | _ -> mpv, None) - | _ -> mpv, None - in Def (na, t, Option.default (Term (CAst.make @@ CHole (None, IntroAnonymous, None))) ty) } + match mpv with + | Term t -> (match t with + | { CAst.v = CCast (t, (CastConv ty | CastVM ty | CastNative ty)) } -> Term t, Some (Term ty) + | _ -> mpv, None) + | _ -> mpv, None + in Def (na, t, Option.default (Term (CAst.make @@ CHole (None, IntroAnonymous, None))) ty) } ] ] ; match_context_rule: @@ -337,9 +337,9 @@ GRAMMAR EXTEND Gram | g = OPT toplevel_selector; "{" -> { Vernacexpr.VernacSubproof g } ] ] ; command: - [ [ IDENT "Proof"; "with"; ta = Pltac.tactic; + [ [ IDENT "Proof"; "with"; ta = Pltac.tactic; l = OPT [ "using"; l = G_vernac.section_subset_expr -> { l } ] -> - { Vernacexpr.VernacProof (Some (in_tac ta), l) } + { Vernacexpr.VernacProof (Some (in_tac ta), l) } | IDENT "Proof"; "using"; l = G_vernac.section_subset_expr; ta = OPT [ "with"; ta = Pltac.tactic -> { in_tac ta } ] -> { Vernacexpr.VernacProof (ta,Some l) } ] ] diff --git a/plugins/ltac/g_obligations.mlg b/plugins/ltac/g_obligations.mlg index 61cc77c42a..5a7a634ed0 100644 --- a/plugins/ltac/g_obligations.mlg +++ b/plugins/ltac/g_obligations.mlg @@ -20,7 +20,7 @@ open Stdarg open Tacarg open Extraargs -let (set_default_tactic, get_default_tactic, print_default_tactic) = +let (set_default_tactic, get_default_tactic, print_default_tactic) = Tactic_option.declare_tactic_option "Program tactic" let () = diff --git a/plugins/ltac/g_rewrite.mlg b/plugins/ltac/g_rewrite.mlg index d25448b5cb..2209edcbb4 100644 --- a/plugins/ltac/g_rewrite.mlg +++ b/plugins/ltac/g_rewrite.mlg @@ -71,7 +71,7 @@ END type raw_strategy = (constr_expr, Tacexpr.raw_red_expr) strategy_ast type glob_strategy = (glob_constr_and_expr, Tacexpr.raw_red_expr) strategy_ast -let interp_strategy ist gl s = +let interp_strategy ist gl s = let sigma = project gl in sigma, strategy_of_ast s let glob_strategy ist s = map_strategy (Tacintern.intern_constr ist) (fun c -> c) s diff --git a/plugins/ltac/g_tactic.mlg b/plugins/ltac/g_tactic.mlg index eb282d1f83..d82eadcfc7 100644 --- a/plugins/ltac/g_tactic.mlg +++ b/plugins/ltac/g_tactic.mlg @@ -45,10 +45,10 @@ let test_lpar_id_coloneq = (match stream_nth 1 strm with | IDENT _ -> (match stream_nth 2 strm with - | KEYWORD ":=" -> () - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) + | KEYWORD ":=" -> () + | _ -> err ()) + | _ -> err ()) + | _ -> err ()) (* Hack to recognize "(x)" *) let test_lpar_id_rpar = @@ -59,10 +59,10 @@ let test_lpar_id_rpar = (match stream_nth 1 strm with | IDENT _ -> (match stream_nth 2 strm with - | KEYWORD ")" -> () - | _ -> err ()) - | _ -> err ()) - | _ -> err ()) + | KEYWORD ")" -> () + | _ -> err ()) + | _ -> err ()) + | _ -> err ()) (* idem for (x:=t) and (1:=t) *) let test_lpar_idnum_coloneq = @@ -86,22 +86,22 @@ let check_for_coloneq = Pcoq.Entry.of_parser "lpar_id_colon" (fun _ strm -> let rec skip_to_rpar p n = - match List.last (Stream.npeek n strm) with + match List.last (Stream.npeek n strm) with | KEYWORD "(" -> skip_to_rpar (p+1) (n+1) | KEYWORD ")" -> if Int.equal p 0 then n+1 else skip_to_rpar (p-1) (n+1) - | KEYWORD "." -> err () - | _ -> skip_to_rpar p (n+1) in + | KEYWORD "." -> err () + | _ -> skip_to_rpar p (n+1) in let rec skip_names n = - match List.last (Stream.npeek n strm) with + match List.last (Stream.npeek n strm) with | IDENT _ | KEYWORD "_" -> skip_names (n+1) - | KEYWORD ":" -> skip_to_rpar 0 (n+1) (* skip a constr *) - | _ -> err () in + | KEYWORD ":" -> skip_to_rpar 0 (n+1) (* skip a constr *) + | _ -> err () in let rec skip_binders n = - match List.last (Stream.npeek n strm) with + match List.last (Stream.npeek n strm) with | KEYWORD "(" -> skip_binders (skip_names (n+1)) | IDENT _ | KEYWORD "_" -> skip_binders (n+1) - | KEYWORD ":=" -> () - | _ -> err () in + | KEYWORD ":=" -> () + | _ -> err () in match stream_nth 0 strm with | KEYWORD "(" -> skip_binders 2 | _ -> err ()) @@ -110,8 +110,8 @@ let lookup_at_as_comma = Pcoq.Entry.of_parser "lookup_at_as_comma" (fun _ strm -> match stream_nth 0 strm with - | KEYWORD (","|"at"|"as") -> () - | _ -> err ()) + | KEYWORD (","|"at"|"as") -> () + | _ -> err ()) open Constr open Prim @@ -164,7 +164,7 @@ let mkTacCase with_evar = function | ic -> if List.exists (function ((_, ElimOnAnonHyp _),_,_) -> true | _ -> false) (fst ic) then - user_err Pp.(str "Use of numbers as direct arguments of 'case' is not supported."); + user_err Pp.(str "Use of numbers as direct arguments of 'case' is not supported."); TacInductionDestruct (false,with_evar,ic) let rec mkCLambdaN_simple_loc ?loc bll c = @@ -188,7 +188,7 @@ let merge_occurrences loc cl = function | None -> if Locusops.clause_with_generic_occurrences cl then (None, cl) else - user_err ~loc (str "Found an \"at\" clause without \"with\" clause.") + user_err ~loc (str "Found an \"at\" clause without \"with\" clause.") | Some (occs, p) -> let ans = match occs with | AllOccurrences -> cl @@ -264,7 +264,7 @@ GRAMMAR EXTEND Gram occs_nums: [ [ nl = LIST1 nat_or_var -> { OnlyOccurrences nl } | "-"; n = nat_or_var; nl = LIST0 int_or_var -> - (* have used int_or_var instead of nat_or_var for compatibility *) + (* have used int_or_var instead of nat_or_var for compatibility *) { AllOccurrencesBut (List.map (Locusops.or_var_map abs) (n::nl)) } ] ] ; occs: @@ -296,12 +296,12 @@ GRAMMAR EXTEND Gram tc = LIST1 simple_intropattern SEP "," ; ")" -> { IntroAndPattern (si::tc) } | "("; si = simple_intropattern; "&"; - tc = LIST1 simple_intropattern SEP "&" ; ")" -> - (* (A & B & C) is translated into (A,(B,C)) *) - { let rec pairify = function - | ([]|[_]|[_;_]) as l -> l + tc = LIST1 simple_intropattern SEP "&" ; ")" -> + (* (A & B & C) is translated into (A,(B,C)) *) + { let rec pairify = function + | ([]|[_]|[_;_]) as l -> l | t::q -> [t; CAst.make ?loc:(loc_of_ne_list q) (IntroAction (IntroOrAndPattern (IntroAndPattern (pairify q))))] - in IntroAndPattern (pairify (si::tc)) } ] ] + in IntroAndPattern (pairify (si::tc)) } ] ] ; equality_intropattern: [ [ "->" -> { IntroRewrite true } @@ -550,24 +550,24 @@ GRAMMAR EXTEND Gram | IDENT "case"; icl = induction_clause_list -> { TacAtom (CAst.make ~loc @@ mkTacCase false icl) } | IDENT "ecase"; icl = induction_clause_list -> { TacAtom (CAst.make ~loc @@ mkTacCase true icl) } | "fix"; id = ident; n = natural; "with"; fd = LIST1 fixdecl -> - { TacAtom (CAst.make ~loc @@ TacMutualFix (id,n,List.map mk_fix_tac fd)) } + { TacAtom (CAst.make ~loc @@ TacMutualFix (id,n,List.map mk_fix_tac fd)) } | "cofix"; id = ident; "with"; fd = LIST1 cofixdecl -> - { TacAtom (CAst.make ~loc @@ TacMutualCofix (id,List.map mk_cofix_tac fd)) } + { TacAtom (CAst.make ~loc @@ TacMutualCofix (id,List.map mk_cofix_tac fd)) } | IDENT "pose"; bl = bindings_with_parameters -> - { let (id,b) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (false,Names.Name.Name id,b,Locusops.nowhere,true,None)) } + { let (id,b) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (false,Names.Name.Name id,b,Locusops.nowhere,true,None)) } | IDENT "pose"; b = constr; na = as_name -> - { TacAtom (CAst.make ~loc @@ TacLetTac (false,na,b,Locusops.nowhere,true,None)) } + { TacAtom (CAst.make ~loc @@ TacLetTac (false,na,b,Locusops.nowhere,true,None)) } | IDENT "epose"; bl = bindings_with_parameters -> - { let (id,b) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (true,Names.Name id,b,Locusops.nowhere,true,None)) } + { let (id,b) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (true,Names.Name id,b,Locusops.nowhere,true,None)) } | IDENT "epose"; b = constr; na = as_name -> - { TacAtom (CAst.make ~loc @@ TacLetTac (true,na,b,Locusops.nowhere,true,None)) } + { TacAtom (CAst.make ~loc @@ TacLetTac (true,na,b,Locusops.nowhere,true,None)) } | IDENT "set"; bl = bindings_with_parameters; p = clause_dft_concl -> - { let (id,c) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (false,Names.Name.Name id,c,p,true,None)) } + { let (id,c) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (false,Names.Name.Name id,c,p,true,None)) } | IDENT "set"; c = constr; na = as_name; p = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacLetTac (false,na,c,p,true,None)) } | IDENT "eset"; bl = bindings_with_parameters; p = clause_dft_concl -> - { let (id,c) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (true,Names.Name id,c,p,true,None)) } + { let (id,c) = bl in TacAtom (CAst.make ~loc @@ TacLetTac (true,Names.Name id,c,p,true,None)) } | IDENT "eset"; c = constr; na = as_name; p = clause_dft_concl -> { TacAtom (CAst.make ~loc @@ TacLetTac (true,na,c,p,true,None)) } | IDENT "remember"; c = constr; na = as_name; e = eqn_ipat; @@ -579,51 +579,51 @@ GRAMMAR EXTEND Gram (* Alternative syntax for "pose proof c as id" *) | IDENT "assert"; test_lpar_id_coloneq; "("; lid = identref; ":="; - c = lconstr; ")" -> + c = lconstr; ")" -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (false,true,None,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } | IDENT "eassert"; test_lpar_id_coloneq; "("; lid = identref; ":="; - c = lconstr; ")" -> + c = lconstr; ")" -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (true,true,None,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } (* Alternative syntax for "assert c as id by tac" *) | IDENT "assert"; test_lpar_id_colon; "("; lid = identref; ":"; - c = lconstr; ")"; tac=by_tactic -> + c = lconstr; ")"; tac=by_tactic -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (false,true,Some tac,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } | IDENT "eassert"; test_lpar_id_colon; "("; lid = identref; ":"; - c = lconstr; ")"; tac=by_tactic -> + c = lconstr; ")"; tac=by_tactic -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (true,true,Some tac,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } (* Alternative syntax for "enough c as id by tac" *) | IDENT "enough"; test_lpar_id_colon; "("; lid = identref; ":"; - c = lconstr; ")"; tac=by_tactic -> + c = lconstr; ")"; tac=by_tactic -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (false,false,Some tac,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } | IDENT "eenough"; test_lpar_id_colon; "("; lid = identref; ":"; - c = lconstr; ")"; tac=by_tactic -> + c = lconstr; ")"; tac=by_tactic -> { let { CAst.loc = loc; v = id } = lid in TacAtom (CAst.make ?loc @@ TacAssert (true,false,Some tac,Some (CAst.make ?loc @@ IntroNaming (IntroIdentifier id)),c)) } | IDENT "assert"; c = constr; ipat = as_ipat; tac = by_tactic -> - { TacAtom (CAst.make ~loc @@ TacAssert (false,true,Some tac,ipat,c)) } + { TacAtom (CAst.make ~loc @@ TacAssert (false,true,Some tac,ipat,c)) } | IDENT "eassert"; c = constr; ipat = as_ipat; tac = by_tactic -> - { TacAtom (CAst.make ~loc @@ TacAssert (true,true,Some tac,ipat,c)) } + { TacAtom (CAst.make ~loc @@ TacAssert (true,true,Some tac,ipat,c)) } | IDENT "pose"; IDENT "proof"; c = lconstr; ipat = as_ipat -> - { TacAtom (CAst.make ~loc @@ TacAssert (false,true,None,ipat,c)) } + { TacAtom (CAst.make ~loc @@ TacAssert (false,true,None,ipat,c)) } | IDENT "epose"; IDENT "proof"; c = lconstr; ipat = as_ipat -> - { TacAtom (CAst.make ~loc @@ TacAssert (true,true,None,ipat,c)) } + { TacAtom (CAst.make ~loc @@ TacAssert (true,true,None,ipat,c)) } | IDENT "enough"; c = constr; ipat = as_ipat; tac = by_tactic -> - { TacAtom (CAst.make ~loc @@ TacAssert (false,false,Some tac,ipat,c)) } + { TacAtom (CAst.make ~loc @@ TacAssert (false,false,Some tac,ipat,c)) } | IDENT "eenough"; c = constr; ipat = as_ipat; tac = by_tactic -> - { TacAtom (CAst.make ~loc @@ TacAssert (true,false,Some tac,ipat,c)) } + { TacAtom (CAst.make ~loc @@ TacAssert (true,false,Some tac,ipat,c)) } | IDENT "generalize"; c = constr -> - { TacAtom (CAst.make ~loc @@ TacGeneralize [((AllOccurrences,c),Names.Name.Anonymous)]) } + { TacAtom (CAst.make ~loc @@ TacGeneralize [((AllOccurrences,c),Names.Name.Anonymous)]) } | IDENT "generalize"; c = constr; l = LIST1 constr -> - { let gen_everywhere c = ((AllOccurrences,c),Names.Name.Anonymous) in + { let gen_everywhere c = ((AllOccurrences,c),Names.Name.Anonymous) in TacAtom (CAst.make ~loc @@ TacGeneralize (List.map gen_everywhere (c::l))) } | IDENT "generalize"; c = constr; lookup_at_as_comma; nl = occs; na = as_name; @@ -632,41 +632,41 @@ GRAMMAR EXTEND Gram (* Derived basic tactics *) | IDENT "induction"; ic = induction_clause_list -> - { TacAtom (CAst.make ~loc @@ TacInductionDestruct (true,false,ic)) } + { TacAtom (CAst.make ~loc @@ TacInductionDestruct (true,false,ic)) } | IDENT "einduction"; ic = induction_clause_list -> - { TacAtom (CAst.make ~loc @@ TacInductionDestruct(true,true,ic)) } + { TacAtom (CAst.make ~loc @@ TacInductionDestruct(true,true,ic)) } | IDENT "destruct"; icl = induction_clause_list -> - { TacAtom (CAst.make ~loc @@ TacInductionDestruct(false,false,icl)) } + { TacAtom (CAst.make ~loc @@ TacInductionDestruct(false,false,icl)) } | IDENT "edestruct"; icl = induction_clause_list -> - { TacAtom (CAst.make ~loc @@ TacInductionDestruct(false,true,icl)) } + { TacAtom (CAst.make ~loc @@ TacInductionDestruct(false,true,icl)) } (* Equality and inversion *) | IDENT "rewrite"; l = LIST1 oriented_rewriter SEP ","; - cl = clause_dft_concl; t=by_tactic -> { TacAtom (CAst.make ~loc @@ TacRewrite (false,l,cl,t)) } + cl = clause_dft_concl; t=by_tactic -> { TacAtom (CAst.make ~loc @@ TacRewrite (false,l,cl,t)) } | IDENT "erewrite"; l = LIST1 oriented_rewriter SEP ","; - cl = clause_dft_concl; t=by_tactic -> { TacAtom (CAst.make ~loc @@ TacRewrite (true,l,cl,t)) } + cl = clause_dft_concl; t=by_tactic -> { TacAtom (CAst.make ~loc @@ TacRewrite (true,l,cl,t)) } | IDENT "dependent"; k = - [ IDENT "simple"; IDENT "inversion" -> { SimpleInversion } - | IDENT "inversion" -> { FullInversion } - | IDENT "inversion_clear" -> { FullInversionClear } ]; - hyp = quantified_hypothesis; - ids = as_or_and_ipat; co = OPT ["with"; c = constr -> { c } ] -> - { TacAtom (CAst.make ~loc @@ TacInversion (DepInversion (k,co,ids),hyp)) } + [ IDENT "simple"; IDENT "inversion" -> { SimpleInversion } + | IDENT "inversion" -> { FullInversion } + | IDENT "inversion_clear" -> { FullInversionClear } ]; + hyp = quantified_hypothesis; + ids = as_or_and_ipat; co = OPT ["with"; c = constr -> { c } ] -> + { TacAtom (CAst.make ~loc @@ TacInversion (DepInversion (k,co,ids),hyp)) } | IDENT "simple"; IDENT "inversion"; - hyp = quantified_hypothesis; ids = as_or_and_ipat; - cl = in_hyp_list -> - { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (SimpleInversion, cl, ids), hyp)) } + hyp = quantified_hypothesis; ids = as_or_and_ipat; + cl = in_hyp_list -> + { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (SimpleInversion, cl, ids), hyp)) } | IDENT "inversion"; - hyp = quantified_hypothesis; ids = as_or_and_ipat; - cl = in_hyp_list -> - { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (FullInversion, cl, ids), hyp)) } + hyp = quantified_hypothesis; ids = as_or_and_ipat; + cl = in_hyp_list -> + { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (FullInversion, cl, ids), hyp)) } | IDENT "inversion_clear"; - hyp = quantified_hypothesis; ids = as_or_and_ipat; - cl = in_hyp_list -> - { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (FullInversionClear, cl, ids), hyp)) } + hyp = quantified_hypothesis; ids = as_or_and_ipat; + cl = in_hyp_list -> + { TacAtom (CAst.make ~loc @@ TacInversion (NonDepInversion (FullInversionClear, cl, ids), hyp)) } | IDENT "inversion"; hyp = quantified_hypothesis; - "using"; c = constr; cl = in_hyp_list -> - { TacAtom (CAst.make ~loc @@ TacInversion (InversionUsing (c,cl), hyp)) } + "using"; c = constr; cl = in_hyp_list -> + { TacAtom (CAst.make ~loc @@ TacInversion (InversionUsing (c,cl), hyp)) } (* Conversion *) | IDENT "red"; cl = clause_dft_concl -> @@ -696,8 +696,8 @@ GRAMMAR EXTEND Gram (* Change ne doit pas s'appliquer dans un Definition t := Eval ... *) | IDENT "change"; c = conversion; cl = clause_dft_concl -> - { let (oc, c) = c in - let p,cl = merge_occurrences loc cl oc in + { let (oc, c) = c in + let p,cl = merge_occurrences loc cl oc in TacAtom (CAst.make ~loc @@ TacChange (true,p,c,cl)) } | IDENT "change_no_check"; c = conversion; cl = clause_dft_concl -> { let (oc, c) = c in diff --git a/plugins/ltac/pltac.ml b/plugins/ltac/pltac.ml index e3042dc3cb..0e21115474 100644 --- a/plugins/ltac/pltac.ml +++ b/plugins/ltac/pltac.ml @@ -34,7 +34,7 @@ let int_or_var = make_gen_entry utactic "int_or_var" let simple_intropattern = make_gen_entry utactic "simple_intropattern" let in_clause = make_gen_entry utactic "in_clause" -let clause_dft_concl = +let clause_dft_concl = make_gen_entry utactic "clause" diff --git a/plugins/ltac/rewrite.ml b/plugins/ltac/rewrite.ml index 6be358be21..5618fd7bc3 100644 --- a/plugins/ltac/rewrite.ml +++ b/plugins/ltac/rewrite.ml @@ -76,7 +76,7 @@ let find_global dir s = let gr = lazy (find_reference dir s) in fun (evd,cstrs) -> let (evd, c) = Evarutil.new_global evd (Lazy.force gr) in - (evd, cstrs), c + (evd, cstrs), c (** Utility for dealing with polymorphic applications *) @@ -122,7 +122,7 @@ let app_poly_nocheck env evars f args = let app_poly_sort b = if b then app_poly_nocheck else app_poly_check - + let find_class_proof proof_type proof_method env evars carrier relation = try let evars, goal = app_poly_check env evars proof_type [| carrier ; relation |] in @@ -130,7 +130,7 @@ let find_class_proof proof_type proof_method env evars carrier relation = if extends_undefined (goalevars evars) evars' then raise Not_found else app_poly_check env (evars',cstrevars evars) proof_method [| carrier; relation; c |] with e when Logic.catchable_exception e -> raise Not_found - + (** Utility functions *) module GlobalBindings (M : sig @@ -146,7 +146,7 @@ end) = struct let reflexive_type = find_global relation_classes "Reflexive" let reflexive_proof = find_global relation_classes "reflexivity" - + let symmetric_type = find_global relation_classes "Symmetric" let symmetric_proof = find_global relation_classes "symmetry" @@ -201,53 +201,53 @@ end) = struct let get_symmetric_proof env = find_class_proof symmetric_type symmetric_proof env let get_transitive_proof env = find_class_proof transitive_type transitive_proof env - let mk_relation env evd a = + let mk_relation env evd a = app_poly env evd relation [| a |] (** Build an inferred signature from constraints on the arguments and expected output relation *) - + let build_signature evars env m (cstrs : (types * types option) option list) (finalcstr : (types * types option) option) = let mk_relty evars newenv ty obj = match obj with | None | Some (_, None) -> - let evars, relty = mk_relation env evars ty in - if closed0 (goalevars evars) ty then - let env' = Environ.reset_with_named_context (Environ.named_context_val env) env in - new_cstr_evar evars env' relty - else new_cstr_evar evars newenv relty + let evars, relty = mk_relation env evars ty in + if closed0 (goalevars evars) ty then + let env' = Environ.reset_with_named_context (Environ.named_context_val env) env in + new_cstr_evar evars env' relty + else new_cstr_evar evars newenv relty | Some (x, Some rel) -> evars, rel in let rec aux env evars ty l = let t = Reductionops.whd_all env (goalevars evars) ty in - match EConstr.kind (goalevars evars) t, l with + match EConstr.kind (goalevars evars) t, l with | Prod (na, ty, b), obj :: cstrs -> let b = Reductionops.nf_betaiota env (goalevars evars) b in if noccurn (goalevars evars) 1 b (* non-dependent product *) then let ty = Reductionops.nf_betaiota env (goalevars evars) ty in - let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in - let evars, relty = mk_relty evars env ty obj in - let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in + let (evars, b', arg, cstrs) = aux env evars (subst1 mkProp b) cstrs in + let evars, relty = mk_relty evars env ty obj in + let evars, newarg = app_poly env evars respectful [| ty ; b' ; relty ; arg |] in evars, mkProd(na, ty, b), newarg, (ty, Some relty) :: cstrs - else - let (evars, b, arg, cstrs) = + else + let (evars, b, arg, cstrs) = aux (push_rel (LocalAssum (na, ty)) env) evars b cstrs - in + in let ty = Reductionops.nf_betaiota env (goalevars evars) ty in let pred = mkLambda (na, ty, b) in let liftarg = mkLambda (na, ty, arg) in let evars, arg' = app_poly env evars forall_relation [| ty ; pred ; liftarg |] in if Option.is_empty obj then evars, mkProd(na, ty, b), arg', (ty, None) :: cstrs - else user_err Pp.(str "build_signature: no constraint can apply on a dependent argument") - | _, obj :: _ -> anomaly ~label:"build_signature" (Pp.str "not enough products.") - | _, [] -> - (match finalcstr with - | None | Some (_, None) -> + else user_err Pp.(str "build_signature: no constraint can apply on a dependent argument") + | _, obj :: _ -> anomaly ~label:"build_signature" (Pp.str "not enough products.") + | _, [] -> + (match finalcstr with + | None | Some (_, None) -> let t = Reductionops.nf_betaiota env (fst evars) ty in - let evars, rel = mk_relty evars env t None in - evars, t, rel, [t, Some rel] - | Some (t, Some rel) -> evars, t, rel, [t, Some rel]) + let evars, rel = mk_relty evars env t None in + evars, t, rel, [t, Some rel] + | Some (t, Some rel) -> evars, t, rel, [t, Some rel]) in aux env evars m cstrs (** Folding/unfolding of the tactic constants. *) @@ -278,30 +278,30 @@ end) = struct let ap = is_Prop (goalevars evd) ta and bp = is_Prop (goalevars evd) tb in if ap && bp then app_poly env evd impl [| a; b |], unfold_impl else if ap then (* Domain in Prop, CoDomain in Type *) - (app_poly env evd arrow [| a; b |]), unfold_impl - (* (evd, mkProd (Anonymous, a, b)), (fun x -> x) *) + (app_poly env evd arrow [| a; b |]), unfold_impl + (* (evd, mkProd (Anonymous, a, b)), (fun x -> x) *) else if bp then (* Dummy forall *) (app_poly env evd coq_all [| a; mkLambda (make_annot Anonymous Sorts.Relevant, a, lift 1 b) |]), unfold_forall else (* None in Prop, use arrow *) - (app_poly env evd arrow [| a; b |]), unfold_impl + (app_poly env evd arrow [| a; b |]), unfold_impl let rec decomp_pointwise sigma n c = if Int.equal n 0 then c else match EConstr.kind sigma c with | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f -> - decomp_pointwise sigma (pred n) relb + decomp_pointwise sigma (pred n) relb | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f -> - decomp_pointwise sigma (pred n) (Reductionops.beta_applist sigma (arelb, [mkRel 1])) + decomp_pointwise sigma (pred n) (Reductionops.beta_applist sigma (arelb, [mkRel 1])) | _ -> invalid_arg "decomp_pointwise" let rec apply_pointwise sigma rel = function | arg :: args -> (match EConstr.kind sigma rel with | App (f, [| a; b; relb |]) when Termops.is_global sigma (pointwise_relation_ref ()) f -> - apply_pointwise sigma relb args + apply_pointwise sigma relb args | App (f, [| a; b; arelb |]) when Termops.is_global sigma (forall_relation_ref ()) f -> - apply_pointwise sigma (Reductionops.beta_applist sigma (arelb, [arg])) args + apply_pointwise sigma (Reductionops.beta_applist sigma (arelb, [arg])) args | _ -> invalid_arg "apply_pointwise") | [] -> rel @@ -316,36 +316,36 @@ end) = struct let lift_cstr env evars (args : constr list) c ty cstr = let start evars env car = match cstr with - | None | Some (_, None) -> - let evars, rel = mk_relation env evars car in - new_cstr_evar evars env rel + | None | Some (_, None) -> + let evars, rel = mk_relation env evars car in + new_cstr_evar evars env rel | Some (ty, Some rel) -> evars, rel in - let rec aux evars env prod n = + let rec aux evars env prod n = if Int.equal n 0 then start evars env prod else let sigma = goalevars evars in - match EConstr.kind sigma (Reductionops.whd_all env sigma prod) with + match EConstr.kind sigma (Reductionops.whd_all env sigma prod) with | Prod (na, ty, b) -> - if noccurn sigma 1 b then - let b' = lift (-1) b in - let evars, rb = aux evars env b' (pred n) in - app_poly env evars pointwise_relation [| ty; b'; rb |] - else + if noccurn sigma 1 b then + let b' = lift (-1) b in + let evars, rb = aux evars env b' (pred n) in + app_poly env evars pointwise_relation [| ty; b'; rb |] + else let evars, rb = aux evars (push_rel (LocalAssum (na, ty)) env) b (pred n) in - app_poly env evars forall_relation + app_poly env evars forall_relation [| ty; mkLambda (na, ty, b); mkLambda (na, ty, rb) |] - | _ -> raise Not_found - in + | _ -> raise Not_found + in let rec find env c ty = function | [] -> None | arg :: args -> - try let evars, found = aux evars env ty (succ (List.length args)) in - Some (evars, found, c, ty, arg :: args) - with Not_found -> + try let evars, found = aux evars env ty (succ (List.length args)) in + Some (evars, found, c, ty, arg :: args) + with Not_found -> let sigma = goalevars evars in - let ty = Reductionops.whd_all env sigma ty in - find env (mkApp (c, [| arg |])) (prod_applist sigma ty [arg]) args + let ty = Reductionops.whd_all env sigma ty in + find env (mkApp (c, [| arg |])) (prod_applist sigma ty [arg]) args in find env c ty args let unlift_cstr env sigma = function @@ -357,18 +357,18 @@ end) = struct match EConstr.kind sigma t with | App (c, args) when Array.length args >= 2 -> let head = if isApp sigma c then fst (destApp sigma c) else c in - if Termops.is_global sigma (coq_eq_ref ()) head then None - else - (try - let params, args = Array.chop (Array.length args - 2) args in - let env' = push_rel_context rels env in - let (evars, (evar, _)) = Evarutil.new_type_evar env' sigma Evd.univ_flexible in - let evars, inst = - app_poly env (evars,Evar.Set.empty) - rewrite_relation_class [| evar; mkApp (c, params) |] in - let _ = Typeclasses.resolve_one_typeclass env' (goalevars evars) inst in - Some (it_mkProd_or_LetIn t rels) - with e when CErrors.noncritical e -> None) + if Termops.is_global sigma (coq_eq_ref ()) head then None + else + (try + let params, args = Array.chop (Array.length args - 2) args in + let env' = push_rel_context rels env in + let (evars, (evar, _)) = Evarutil.new_type_evar env' sigma Evd.univ_flexible in + let evars, inst = + app_poly env (evars,Evar.Set.empty) + rewrite_relation_class [| evar; mkApp (c, params) |] in + let _ = Typeclasses.resolve_one_typeclass env' (goalevars evars) inst in + Some (it_mkProd_or_LetIn t rels) + with e when CErrors.noncritical e -> None) | _ -> None @@ -386,7 +386,7 @@ let type_app_poly env env evd f args = module PropGlobal = struct module Consts = - struct + struct let relation_classes = ["Coq"; "Classes"; "RelationClasses"] let morphisms = ["Coq"; "Classes"; "Morphisms"] let relation = ["Coq"; "Relations";"Relation_Definitions"], "relation" @@ -399,15 +399,15 @@ module PropGlobal = struct include G include Consts - let inverse env evd car rel = + let inverse env evd car rel = type_app_poly env env evd coq_inverse [| car ; car; mkProp; rel |] (* app_poly env evd coq_inverse [| car ; car; mkProp; rel |] *) end module TypeGlobal = struct - module Consts = - struct + module Consts = + struct let relation_classes = ["Coq"; "Classes"; "CRelationClasses"] let morphisms = ["Coq"; "Classes"; "CMorphisms"] let relation = relation_classes, "crelation" @@ -421,7 +421,7 @@ module TypeGlobal = struct include Consts - let inverse env (evd,cstrs) car rel = + let inverse env (evd,cstrs) car rel = let (evd, sort) = Evarutil.new_Type ~rigid:Evd.univ_flexible evd in app_poly_check env (evd,cstrs) coq_inverse [| car ; car; sort; rel |] @@ -471,12 +471,12 @@ type hypinfo = { holes : Clenv.hole list; } -let get_symmetric_proof b = +let get_symmetric_proof b = if b then PropGlobal.get_symmetric_proof else TypeGlobal.get_symmetric_proof let error_no_relation () = user_err Pp.(str "Cannot find a relation to rewrite.") -let rec decompose_app_rel env evd t = +let rec decompose_app_rel env evd t = (* Head normalize for compatibility with the old meta mechanism *) let t = Reductionops.whd_betaiota evd t in match EConstr.kind evd t with @@ -525,10 +525,10 @@ let decompose_applied_relation env sigma (c,l) = match find_rel ctype with | Some c -> c | None -> - let ctx,t' = Reductionops.splay_prod env sigma ctype in (* Search for underlying eq *) + let ctx,t' = Reductionops.splay_prod env sigma ctype in (* Search for underlying eq *) match find_rel (it_mkProd_or_LetIn t' (List.map (fun (n,t) -> LocalAssum (n, t)) ctx)) with - | Some c -> c - | None -> user_err Pp.(str "Cannot find an homogeneous relation to rewrite.") + | Some c -> c + | None -> user_err Pp.(str "Cannot find an homogeneous relation to rewrite.") let rewrite_db = "rewrite" @@ -644,13 +644,13 @@ let solve_remaining_by env sigma holes by = in List.fold_left solve sigma indep -let no_constraints cstrs = +let no_constraints cstrs = fun ev _ -> not (Evar.Set.mem ev cstrs) let poly_inverse sort = if sort then PropGlobal.inverse else TypeGlobal.inverse -type rewrite_proof = +type rewrite_proof = | RewPrf of constr * constr (** A Relation (R : rew_car -> rew_car -> Prop) and a proof of R rew_from rew_to *) @@ -675,13 +675,13 @@ type rewrite_result = | Success of rewrite_result_info type 'a strategy_input = { state : 'a ; (* a parameter: for instance, a state *) - env : Environ.env ; - unfresh : Id.Set.t; (* Unfresh names *) - term1 : constr ; - ty1 : types ; (* first term and its type (convertible to rew_from) *) - cstr : (bool (* prop *) * constr option) ; - evars : evars } - + env : Environ.env ; + unfresh : Id.Set.t; (* Unfresh names *) + term1 : constr ; + ty1 : types ; (* first term and its type (convertible to rew_from) *) + cstr : (bool (* prop *) * constr option) ; + evars : evars } + type 'a pure_strategy = { strategy : 'a strategy_input -> 'a * rewrite_result (* the updated state and the "result" *) } @@ -723,7 +723,7 @@ let unify_eqn (car, rel, prf, c1, c2, holes, sort) l2r flags env (sigma, cstrs) let rew = { rew_evars; rew_prf; rew_car; rew_from = c1; rew_to = c2; } in let rew = if l2r then rew else symmetry env sort rew in Some rew - with + with | e when Class_tactics.catchable e -> None | Reduction.NotConvertible -> None @@ -740,7 +740,7 @@ let unify_abs (car, rel, prf, c1, c2) l2r sort env (sigma, cstrs) t = let rew = { rew_car = car; rew_from = c1; rew_to = c2; rew_prf; rew_evars; } in let rew = if l2r then rew else symmetry env sort rew in Some rew - with + with | e when Class_tactics.catchable e -> None | Reduction.NotConvertible -> None @@ -766,9 +766,9 @@ let get_rew_prf evars r = match r.rew_prf with let evars, eq_refl = make_eq_refl evars in let rel = mkApp (eq, [| r.rew_car |]) in evars, (rel, mkCast (mkApp (eq_refl, [| r.rew_car; r.rew_from |]), - c, mkApp (rel, [| r.rew_from; r.rew_to |]))) + c, mkApp (rel, [| r.rew_from; r.rew_to |]))) -let poly_subrelation sort = +let poly_subrelation sort = if sort then PropGlobal.subrelation else TypeGlobal.subrelation let resolve_subrelation env avoid car rel sort prf rel' res = @@ -778,8 +778,8 @@ let resolve_subrelation env avoid car rel sort prf rel' res = let evars, subrel = new_cstr_evar evars env app in let appsub = mkApp (subrel, [| res.rew_from ; res.rew_to ; prf |]) in { res with - rew_prf = RewPrf (rel', appsub); - rew_evars = evars } + rew_prf = RewPrf (rel', appsub); + rew_evars = evars } let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) evars = let evars, morph_instance, proj, sigargs, m', args, args' = @@ -790,12 +790,12 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev let morphargs', morphobjs' = Array.chop first args' in let appm = mkApp(m, morphargs) in let appmtype = Typing.unsafe_type_of env (goalevars evars) appm in - let cstrs = List.map - (Option.map (fun r -> r.rew_car, get_opt_rew_rel r.rew_prf)) - (Array.to_list morphobjs') + let cstrs = List.map + (Option.map (fun r -> r.rew_car, get_opt_rew_rel r.rew_prf)) + (Array.to_list morphobjs') in (* Desired signature *) - let evars, appmtype', signature, sigargs = + let evars, appmtype', signature, sigargs = if b then PropGlobal.build_signature evars env appmtype cstrs cstr else TypeGlobal.build_signature evars env appmtype cstrs cstr in @@ -803,16 +803,16 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev let cl_args = [| appmtype' ; signature ; appm |] in let evars, app = app_poly_sort b env evars (if b then PropGlobal.proper_type env else TypeGlobal.proper_type env) cl_args in - let env' = - let dosub, appsub = - if b then PropGlobal.do_subrelation, PropGlobal.apply_subrelation - else TypeGlobal.do_subrelation, TypeGlobal.apply_subrelation + let env' = + let dosub, appsub = + if b then PropGlobal.do_subrelation, PropGlobal.apply_subrelation + else TypeGlobal.do_subrelation, TypeGlobal.apply_subrelation in - EConstr.push_named + EConstr.push_named (LocalDef (make_annot (Id.of_string "do_subrelation") Sorts.Relevant, - snd (app_poly_sort b env evars dosub [||]), - snd (app_poly_nocheck env evars appsub [||]))) - env + snd (app_poly_sort b env evars dosub [||]), + snd (app_poly_nocheck env evars appsub [||]))) + env in let evars, morph = new_cstr_evar evars env' app in evars, morph, morph, sigargs, appm, morphobjs, morphobjs' @@ -820,31 +820,31 @@ let resolve_morphism env avoid oldt m ?(fnewt=fun x -> x) args args' (b,cstr) ev let projargs, subst, evars, respars, typeargs = Array.fold_left2 (fun (acc, subst, evars, sigargs, typeargs') x y -> - let (carrier, relation), sigargs = split_head sigargs in - match relation with - | Some relation -> - let carrier = substl subst carrier - and relation = substl subst relation in - (match y with - | None -> - let evars, proof = - (if b then PropGlobal.proper_proof else TypeGlobal.proper_proof) - env evars carrier relation x in - [ proof ; x ; x ] @ acc, subst, evars, sigargs, x :: typeargs' - | Some r -> + let (carrier, relation), sigargs = split_head sigargs in + match relation with + | Some relation -> + let carrier = substl subst carrier + and relation = substl subst relation in + (match y with + | None -> + let evars, proof = + (if b then PropGlobal.proper_proof else TypeGlobal.proper_proof) + env evars carrier relation x in + [ proof ; x ; x ] @ acc, subst, evars, sigargs, x :: typeargs' + | Some r -> let evars, proof = get_rew_prf evars r in - [ snd proof; r.rew_to; x ] @ acc, subst, evars, - sigargs, r.rew_to :: typeargs') - | None -> - if not (Option.is_empty y) then - user_err Pp.(str "Cannot rewrite inside dependent arguments of a function"); - x :: acc, x :: subst, evars, sigargs, x :: typeargs') + [ snd proof; r.rew_to; x ] @ acc, subst, evars, + sigargs, r.rew_to :: typeargs') + | None -> + if not (Option.is_empty y) then + user_err Pp.(str "Cannot rewrite inside dependent arguments of a function"); + x :: acc, x :: subst, evars, sigargs, x :: typeargs') ([], [], evars, sigargs, []) args args' in let proof = applist (proj, List.rev projargs) in let newt = applist (m', List.rev typeargs) in match respars with - [ a, Some r ] -> evars, proof, substl subst a, substl subst r, oldt, fnewt newt + [ a, Some r ] -> evars, proof, substl subst a, substl subst r, oldt, fnewt newt | _ -> assert(false) let apply_constraint env avoid car rel prf cstr res = @@ -852,7 +852,7 @@ let apply_constraint env avoid car rel prf cstr res = | None -> res | Some r -> resolve_subrelation env avoid car rel (fst cstr) prf r res -let coerce env avoid cstr res = +let coerce env avoid cstr res = let evars, (rel, prf) = get_rew_prf res.rew_evars res in let res = { res with rew_evars = evars } in apply_constraint env avoid res.rew_car rel prf cstr res @@ -860,22 +860,22 @@ let coerce env avoid cstr res = let apply_rule unify loccs : int pure_strategy = let (nowhere_except_in,occs) = convert_occs loccs in let is_occ occ = - if nowhere_except_in - then List.mem occ occs - else not (List.mem occ occs) + if nowhere_except_in + then List.mem occ occs + else not (List.mem occ occs) in { strategy = fun { state = occ ; env ; unfresh ; - term1 = t ; ty1 = ty ; cstr ; evars } -> + term1 = t ; ty1 = ty ; cstr ; evars } -> let unif = if isEvar (goalevars evars) t then None else unify env evars t in - match unif with - | None -> (occ, Fail) + match unif with + | None -> (occ, Fail) | Some rew -> - let occ = succ occ in - if not (is_occ occ) then (occ, Fail) - else if Termops.eq_constr (fst rew.rew_evars) t rew.rew_to then (occ, Identity) - else - let res = { rew with rew_car = ty } in - let res = Success (coerce env unfresh cstr res) in + let occ = succ occ in + if not (is_occ occ) then (occ, Fail) + else if Termops.eq_constr (fst rew.rew_evars) t rew.rew_to then (occ, Identity) + else + let res = { rew with rew_car = ty } in + let res = Success (coerce env unfresh cstr res) in (occ, res) } @@ -893,10 +893,10 @@ let apply_lemma l2r flags oc by loccs : strategy = { strategy = | Some rew -> Some rew in let _, res = (apply_rule unify loccs).strategy { input with - state = 0 ; - evars } in + state = 0 ; + evars } in (), res - } + } let e_app_poly env evars f args = let evars', c = app_poly_nocheck env !evars f args in @@ -905,16 +905,16 @@ let e_app_poly env evars f args = let make_leibniz_proof env c ty r = let evars = ref r.rew_evars in - let prf = + let prf = match r.rew_prf with - | RewPrf (rel, prf) -> - let rel = e_app_poly env evars coq_eq [| ty |] in - let prf = - e_app_poly env evars coq_f_equal - [| r.rew_car; ty; + | RewPrf (rel, prf) -> + let rel = e_app_poly env evars coq_eq [| ty |] in + let prf = + e_app_poly env evars coq_f_equal + [| r.rew_car; ty; mkLambda (make_annot Anonymous Sorts.Relevant, r.rew_car, c); - r.rew_from; r.rew_to; prf |] - in RewPrf (rel, prf) + r.rew_from; r.rew_to; prf |] + in RewPrf (rel, prf) | RewCast k -> r.rew_prf in { rew_car = ty; rew_evars = !evars; @@ -923,39 +923,39 @@ let make_leibniz_proof env c ty r = let reset_env env = let env' = Global.env_of_context (Environ.named_context_val env) in Environ.push_rel_context (Environ.rel_context env) env' - + let fold_match ?(force=false) env sigma c = let (ci, p, c, brs) = destCase sigma c in let cty = Retyping.get_type_of env sigma c in - let dep, pred, exists, (sk,eff) = + let dep, pred, exists, (sk,eff) = let env', ctx, body = let ctx, pred = decompose_lam_assum sigma p in let env' = push_rel_context ctx env in - env', ctx, pred + env', ctx, pred in let sortp = Retyping.get_sort_family_of env' sigma body in let sortc = Retyping.get_sort_family_of env sigma cty in let dep = not (noccurn sigma 1 body) in let pred = if dep then p else - it_mkProd_or_LetIn (subst1 mkProp body) (List.tl ctx) + it_mkProd_or_LetIn (subst1 mkProp body) (List.tl ctx) in - let sk = + let sk = if sortp == Sorts.InProp then - if sortc == Sorts.InProp then - if dep then case_dep_scheme_kind_from_prop - else case_scheme_kind_from_prop - else ( + if sortc == Sorts.InProp then + if dep then case_dep_scheme_kind_from_prop + else case_scheme_kind_from_prop + else ( if dep then case_dep_scheme_kind_from_type_in_prop else case_scheme_kind_from_type) else ((* sortc <> InProp by typing *) - if dep - then case_dep_scheme_kind_from_type - else case_scheme_kind_from_type) - in + if dep + then case_dep_scheme_kind_from_type + else case_scheme_kind_from_type) + in let exists = Ind_tables.check_scheme sk ci.ci_ind in if exists || force then - dep, pred, exists, Ind_tables.find_scheme sk ci.ci_ind + dep, pred, exists, Ind_tables.find_scheme sk ci.ci_ind else raise Not_found in let app = @@ -963,7 +963,7 @@ let fold_match ?(force=false) env sigma c = let pars, args = List.chop ci.ci_npar args in let meths = List.map (fun br -> br) (Array.to_list brs) in applist (mkConst sk, pars @ [pred] @ meths @ args @ [c]) - in + in sk, (if exists then env else reset_env env), app, eff let unfold_match env sigma sk app = @@ -971,128 +971,128 @@ let unfold_match env sigma sk app = | App (f', args) when Constant.equal (fst (destConst sigma f')) sk -> let v = Environ.constant_value_in (Global.env ()) (sk,Univ.Instance.empty)(*FIXME*) in let v = EConstr.of_constr v in - Reductionops.whd_beta sigma (mkApp (v, args)) + Reductionops.whd_beta sigma (mkApp (v, args)) | _ -> app let is_rew_cast = function RewCast _ -> true | _ -> false let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = let rec aux { state ; env ; unfresh ; - term1 = t ; ty1 = ty ; cstr = (prop, cstr) ; evars } = + term1 = t ; ty1 = ty ; cstr = (prop, cstr) ; evars } = let cstr' = Option.map (fun c -> (ty, Some c)) cstr in match EConstr.kind (goalevars evars) t with | App (m, args) -> - let rewrite_args state success = - let state, (args', evars', progress) = - Array.fold_left - (fun (state, (acc, evars, progress)) arg -> - if not (Option.is_empty progress) && not all then - state, (None :: acc, evars, progress) - else - let argty = Retyping.get_type_of env (goalevars evars) arg in - let state, res = s.strategy { state ; env ; - unfresh ; - term1 = arg ; ty1 = argty ; - cstr = (prop,None) ; - evars } in - let res' = - match res with - | Identity -> - let progress = if Option.is_empty progress then Some false else progress in - (None :: acc, evars, progress) - | Success r -> - (Some r :: acc, r.rew_evars, Some true) - | Fail -> (None :: acc, evars, progress) - in state, res') - (state, ([], evars, success)) args - in - let res = - match progress with - | None -> Fail - | Some false -> Identity - | Some true -> - let args' = Array.of_list (List.rev args') in - if Array.exists - (function - | None -> false - | Some r -> not (is_rew_cast r.rew_prf)) args' - then - let evars', prf, car, rel, c1, c2 = - resolve_morphism env unfresh t m args args' (prop, cstr') evars' - in - let res = { rew_car = ty; rew_from = c1; - rew_to = c2; rew_prf = RewPrf (rel, prf); - rew_evars = evars' } - in Success res - else - let args' = Array.map2 - (fun aorig anew -> - match anew with None -> aorig - | Some r -> r.rew_to) args args' - in - let res = { rew_car = ty; rew_from = t; - rew_to = mkApp (m, args'); rew_prf = RewCast DEFAULTcast; - rew_evars = evars' } - in Success res - in state, res - in - if flags.on_morphisms then - let mty = Retyping.get_type_of env (goalevars evars) m in - let evars, cstr', m, mty, argsl, args = - let argsl = Array.to_list args in - let lift = if prop then PropGlobal.lift_cstr else TypeGlobal.lift_cstr in - match lift env evars argsl m mty None with - | Some (evars, cstr', m, mty, args) -> - evars, Some cstr', m, mty, args, Array.of_list args - | None -> evars, None, m, mty, argsl, args - in - let state, m' = s.strategy { state ; env ; unfresh ; - term1 = m ; ty1 = mty ; - cstr = (prop, cstr') ; evars } in - match m' with - | Fail -> rewrite_args state None (* Standard path, try rewrite on arguments *) - | Identity -> rewrite_args state (Some false) - | Success r -> - (* We rewrote the function and get a proof of pointwise rel for the arguments. - We just apply it. *) - let prf = match r.rew_prf with - | RewPrf (rel, prf) -> - let app = if prop then PropGlobal.apply_pointwise - else TypeGlobal.apply_pointwise - in - RewPrf (app (goalevars evars) rel argsl, mkApp (prf, args)) - | x -> x - in - let res = - { rew_car = Reductionops.hnf_prod_appvect env (goalevars evars) r.rew_car args; - rew_from = mkApp(r.rew_from, args); rew_to = mkApp(r.rew_to, args); - rew_prf = prf; rew_evars = r.rew_evars } - in - let res = - match prf with - | RewPrf (rel, prf) -> - Success (apply_constraint env unfresh res.rew_car - rel prf (prop,cstr) res) - | _ -> Success res - in state, res - else rewrite_args state None - + let rewrite_args state success = + let state, (args', evars', progress) = + Array.fold_left + (fun (state, (acc, evars, progress)) arg -> + if not (Option.is_empty progress) && not all then + state, (None :: acc, evars, progress) + else + let argty = Retyping.get_type_of env (goalevars evars) arg in + let state, res = s.strategy { state ; env ; + unfresh ; + term1 = arg ; ty1 = argty ; + cstr = (prop,None) ; + evars } in + let res' = + match res with + | Identity -> + let progress = if Option.is_empty progress then Some false else progress in + (None :: acc, evars, progress) + | Success r -> + (Some r :: acc, r.rew_evars, Some true) + | Fail -> (None :: acc, evars, progress) + in state, res') + (state, ([], evars, success)) args + in + let res = + match progress with + | None -> Fail + | Some false -> Identity + | Some true -> + let args' = Array.of_list (List.rev args') in + if Array.exists + (function + | None -> false + | Some r -> not (is_rew_cast r.rew_prf)) args' + then + let evars', prf, car, rel, c1, c2 = + resolve_morphism env unfresh t m args args' (prop, cstr') evars' + in + let res = { rew_car = ty; rew_from = c1; + rew_to = c2; rew_prf = RewPrf (rel, prf); + rew_evars = evars' } + in Success res + else + let args' = Array.map2 + (fun aorig anew -> + match anew with None -> aorig + | Some r -> r.rew_to) args args' + in + let res = { rew_car = ty; rew_from = t; + rew_to = mkApp (m, args'); rew_prf = RewCast DEFAULTcast; + rew_evars = evars' } + in Success res + in state, res + in + if flags.on_morphisms then + let mty = Retyping.get_type_of env (goalevars evars) m in + let evars, cstr', m, mty, argsl, args = + let argsl = Array.to_list args in + let lift = if prop then PropGlobal.lift_cstr else TypeGlobal.lift_cstr in + match lift env evars argsl m mty None with + | Some (evars, cstr', m, mty, args) -> + evars, Some cstr', m, mty, args, Array.of_list args + | None -> evars, None, m, mty, argsl, args + in + let state, m' = s.strategy { state ; env ; unfresh ; + term1 = m ; ty1 = mty ; + cstr = (prop, cstr') ; evars } in + match m' with + | Fail -> rewrite_args state None (* Standard path, try rewrite on arguments *) + | Identity -> rewrite_args state (Some false) + | Success r -> + (* We rewrote the function and get a proof of pointwise rel for the arguments. + We just apply it. *) + let prf = match r.rew_prf with + | RewPrf (rel, prf) -> + let app = if prop then PropGlobal.apply_pointwise + else TypeGlobal.apply_pointwise + in + RewPrf (app (goalevars evars) rel argsl, mkApp (prf, args)) + | x -> x + in + let res = + { rew_car = Reductionops.hnf_prod_appvect env (goalevars evars) r.rew_car args; + rew_from = mkApp(r.rew_from, args); rew_to = mkApp(r.rew_to, args); + rew_prf = prf; rew_evars = r.rew_evars } + in + let res = + match prf with + | RewPrf (rel, prf) -> + Success (apply_constraint env unfresh res.rew_car + rel prf (prop,cstr) res) + | _ -> Success res + in state, res + else rewrite_args state None + | Prod (n, x, b) when noccurn (goalevars evars) 1 b -> - let b = subst1 mkProp b in - let tx = Retyping.get_type_of env (goalevars evars) x - and tb = Retyping.get_type_of env (goalevars evars) b in - let arr = if prop then PropGlobal.arrow_morphism - else TypeGlobal.arrow_morphism - in - let (evars', mor), unfold = arr env evars tx tb x b in - let state, res = aux { state ; env ; unfresh ; - term1 = mor ; ty1 = ty ; - cstr = (prop,cstr) ; evars = evars' } in - let res = - match res with - | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to } - | Fail | Identity -> res - in state, res + let b = subst1 mkProp b in + let tx = Retyping.get_type_of env (goalevars evars) x + and tb = Retyping.get_type_of env (goalevars evars) b in + let arr = if prop then PropGlobal.arrow_morphism + else TypeGlobal.arrow_morphism + in + let (evars', mor), unfold = arr env evars tx tb x b in + let state, res = aux { state ; env ; unfresh ; + term1 = mor ; ty1 = ty ; + cstr = (prop,cstr) ; evars = evars' } in + let res = + match res with + | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to } + | Fail | Identity -> res + in state, res (* if x' = None && flags.under_lambdas then *) (* let lam = mkLambda (n, x, b) in *) @@ -1110,23 +1110,23 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = | Prod (n, dom, codom) -> let lam = mkLambda (n, dom, codom) in - let (evars', app), unfold = - if eq_constr (fst evars) ty mkProp then - (app_poly_sort prop env evars coq_all [| dom; lam |]), TypeGlobal.unfold_all - else - let forall = if prop then PropGlobal.coq_forall else TypeGlobal.coq_forall in - (app_poly_sort prop env evars forall [| dom; lam |]), TypeGlobal.unfold_forall - in - let state, res = aux { state ; env ; unfresh ; - term1 = app ; ty1 = ty ; - cstr = (prop,cstr) ; evars = evars' } in - let res = - match res with - | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to } - | Fail | Identity -> res - in state, res - -(* TODO: real rewriting under binders: introduce x x' (H : R x x') and rewrite with + let (evars', app), unfold = + if eq_constr (fst evars) ty mkProp then + (app_poly_sort prop env evars coq_all [| dom; lam |]), TypeGlobal.unfold_all + else + let forall = if prop then PropGlobal.coq_forall else TypeGlobal.coq_forall in + (app_poly_sort prop env evars forall [| dom; lam |]), TypeGlobal.unfold_forall + in + let state, res = aux { state ; env ; unfresh ; + term1 = app ; ty1 = ty ; + cstr = (prop,cstr) ; evars = evars' } in + let res = + match res with + | Success r -> Success { r with rew_to = unfold (goalevars r.rew_evars) r.rew_to } + | Fail | Identity -> res + in state, res + +(* TODO: real rewriting under binders: introduce x x' (H : R x x') and rewrite with H at any occurrence of x. Ask for (R ==> R') for the lambda. Formalize this. B. Barras' idea is to have a context of relations, of length 1, with Σ for gluing dependent relations and using projections to get them out. @@ -1158,88 +1158,88 @@ let subterm all flags (s : 'a pure_strategy) : 'a pure_strategy = let n' = map_annot (Nameops.Name.map (fun id -> Tactics.fresh_id_in_env unfresh id env)) n in let open Context.Rel.Declaration in let env' = EConstr.push_rel (LocalAssum (n', t)) env in - let bty = Retyping.get_type_of env' (goalevars evars) b in - let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in - let state, b' = s.strategy { state ; env = env' ; unfresh ; - term1 = b ; ty1 = bty ; - cstr = (prop, unlift env evars cstr) ; - evars } in - let res = - match b' with - | Success r -> - let r = match r.rew_prf with - | RewPrf (rel, prf) -> - let point = if prop then PropGlobal.pointwise_or_dep_relation else - TypeGlobal.pointwise_or_dep_relation - in + let bty = Retyping.get_type_of env' (goalevars evars) b in + let unlift = if prop then PropGlobal.unlift_cstr else TypeGlobal.unlift_cstr in + let state, b' = s.strategy { state ; env = env' ; unfresh ; + term1 = b ; ty1 = bty ; + cstr = (prop, unlift env evars cstr) ; + evars } in + let res = + match b' with + | Success r -> + let r = match r.rew_prf with + | RewPrf (rel, prf) -> + let point = if prop then PropGlobal.pointwise_or_dep_relation else + TypeGlobal.pointwise_or_dep_relation + in let evars, rel = point env r.rew_evars n'.binder_name t r.rew_car rel in let prf = mkLambda (n', t, prf) in - { r with rew_prf = RewPrf (rel, prf); rew_evars = evars } - | x -> r - in - Success { r with + { r with rew_prf = RewPrf (rel, prf); rew_evars = evars } + | x -> r + in + Success { r with rew_car = mkProd (n, t, r.rew_car); rew_from = mkLambda(n, t, r.rew_from); rew_to = mkLambda (n, t, r.rew_to) } - | Fail | Identity -> b' - in state, res - + | Fail | Identity -> b' + in state, res + | Case (ci, p, c, brs) -> - let cty = Retyping.get_type_of env (goalevars evars) c in - let evars', eqty = app_poly_sort prop env evars coq_eq [| cty |] in - let cstr' = Some eqty in - let state, c' = s.strategy { state ; env ; unfresh ; - term1 = c ; ty1 = cty ; - cstr = (prop, cstr') ; evars = evars' } in - let state, res = - match c' with - | Success r -> - let case = mkCase (ci, lift 1 p, mkRel 1, Array.map (lift 1) brs) in - let res = make_leibniz_proof env case ty r in - state, Success (coerce env unfresh (prop,cstr) res) - | Fail | Identity -> - if Array.for_all (Int.equal 0) ci.ci_cstr_ndecls then - let evars', eqty = app_poly_sort prop env evars coq_eq [| ty |] in - let cstr = Some eqty in - let state, found, brs' = Array.fold_left - (fun (state, found, acc) br -> - if not (Option.is_empty found) then - (state, found, fun x -> lift 1 br :: acc x) - else - let state, res = s.strategy { state ; env ; unfresh ; - term1 = br ; ty1 = ty ; - cstr = (prop,cstr) ; evars } in - match res with - | Success r -> (state, Some r, fun x -> mkRel 1 :: acc x) - | Fail | Identity -> (state, None, fun x -> lift 1 br :: acc x)) - (state, None, fun x -> []) brs - in - match found with - | Some r -> - let ctxc = mkCase (ci, lift 1 p, lift 1 c, Array.of_list (List.rev (brs' c'))) in - state, Success (make_leibniz_proof env ctxc ty r) - | None -> state, c' - else - match try Some (fold_match env (goalevars evars) t) with Not_found -> None with - | None -> state, c' - | Some (cst, _, t', eff (*FIXME*)) -> - let state, res = aux { state ; env ; unfresh ; - term1 = t' ; ty1 = ty ; - cstr = (prop,cstr) ; evars } in - let res = - match res with - | Success prf -> - Success { prf with - rew_from = t; - rew_to = unfold_match env (goalevars evars) cst prf.rew_to } - | x' -> c' - in state, res - in - let res = - match res with - | Success r -> Success (coerce env unfresh (prop,cstr) r) - | Fail | Identity -> res - in state, res + let cty = Retyping.get_type_of env (goalevars evars) c in + let evars', eqty = app_poly_sort prop env evars coq_eq [| cty |] in + let cstr' = Some eqty in + let state, c' = s.strategy { state ; env ; unfresh ; + term1 = c ; ty1 = cty ; + cstr = (prop, cstr') ; evars = evars' } in + let state, res = + match c' with + | Success r -> + let case = mkCase (ci, lift 1 p, mkRel 1, Array.map (lift 1) brs) in + let res = make_leibniz_proof env case ty r in + state, Success (coerce env unfresh (prop,cstr) res) + | Fail | Identity -> + if Array.for_all (Int.equal 0) ci.ci_cstr_ndecls then + let evars', eqty = app_poly_sort prop env evars coq_eq [| ty |] in + let cstr = Some eqty in + let state, found, brs' = Array.fold_left + (fun (state, found, acc) br -> + if not (Option.is_empty found) then + (state, found, fun x -> lift 1 br :: acc x) + else + let state, res = s.strategy { state ; env ; unfresh ; + term1 = br ; ty1 = ty ; + cstr = (prop,cstr) ; evars } in + match res with + | Success r -> (state, Some r, fun x -> mkRel 1 :: acc x) + | Fail | Identity -> (state, None, fun x -> lift 1 br :: acc x)) + (state, None, fun x -> []) brs + in + match found with + | Some r -> + let ctxc = mkCase (ci, lift 1 p, lift 1 c, Array.of_list (List.rev (brs' c'))) in + state, Success (make_leibniz_proof env ctxc ty r) + | None -> state, c' + else + match try Some (fold_match env (goalevars evars) t) with Not_found -> None with + | None -> state, c' + | Some (cst, _, t', eff (*FIXME*)) -> + let state, res = aux { state ; env ; unfresh ; + term1 = t' ; ty1 = ty ; + cstr = (prop,cstr) ; evars } in + let res = + match res with + | Success prf -> + Success { prf with + rew_from = t; + rew_to = unfold_match env (goalevars evars) cst prf.rew_to } + | x' -> c' + in state, res + in + let res = + match res with + | Success r -> Success (coerce env unfresh (prop,cstr) r) + | Fail | Identity -> res + in state, res | _ -> state, Fail in { strategy = aux } @@ -1249,15 +1249,15 @@ let one_subterm = subterm false default_flags (** Requires transitivity of the rewrite step, if not a reduction. Not tail-recursive. *) -let transitivity state env unfresh prop (res : rewrite_result_info) (next : 'a pure_strategy) : +let transitivity state env unfresh prop (res : rewrite_result_info) (next : 'a pure_strategy) : 'a * rewrite_result = let state, nextres = next.strategy { state ; env ; unfresh ; - term1 = res.rew_to ; ty1 = res.rew_car ; - cstr = (prop, get_opt_rew_rel res.rew_prf) ; - evars = res.rew_evars } - in - let res = + term1 = res.rew_to ; ty1 = res.rew_car ; + cstr = (prop, get_opt_rew_rel res.rew_prf) ; + evars = res.rew_evars } + in + let res = match nextres with | Fail -> Fail | Identity -> Success res @@ -1265,21 +1265,21 @@ let transitivity state env unfresh prop (res : rewrite_result_info) (next : 'a p match res.rew_prf with | RewCast c -> Success { res' with rew_from = res.rew_from } | RewPrf (rew_rel, rew_prf) -> - match res'.rew_prf with - | RewCast _ -> Success { res with rew_to = res'.rew_to } - | RewPrf (res'_rel, res'_prf) -> - let trans = - if prop then PropGlobal.transitive_type - else TypeGlobal.transitive_type - in - let evars, prfty = - app_poly_sort prop env res'.rew_evars trans [| res.rew_car; rew_rel |] - in - let evars, prf = new_cstr_evar evars env prfty in - let prf = mkApp (prf, [|res.rew_from; res'.rew_from; res'.rew_to; - rew_prf; res'_prf |]) - in Success { res' with rew_from = res.rew_from; - rew_evars = evars; rew_prf = RewPrf (res'_rel, prf) } + match res'.rew_prf with + | RewCast _ -> Success { res with rew_to = res'.rew_to } + | RewPrf (res'_rel, res'_prf) -> + let trans = + if prop then PropGlobal.transitive_type + else TypeGlobal.transitive_type + in + let evars, prfty = + app_poly_sort prop env res'.rew_evars trans [| res.rew_car; rew_rel |] + in + let evars, prf = new_cstr_evar evars env prfty in + let prf = mkApp (prf, [|res.rew_from; res'.rew_from; res'.rew_to; + rew_prf; res'_prf |]) + in Success { res' with rew_from = res.rew_from; + rew_evars = evars; rew_prf = RewPrf (res'_rel, prf) } in state, res (** Rewriting strategies. @@ -1299,54 +1299,54 @@ module Strategies = let refl : 'a pure_strategy = { strategy = - fun { state ; env ; - term1 = t ; ty1 = ty ; - cstr = (prop,cstr) ; evars } -> - let evars, rel = match cstr with - | None -> - let mkr = if prop then PropGlobal.mk_relation else TypeGlobal.mk_relation in - let evars, rty = mkr env evars ty in - new_cstr_evar evars env rty - | Some r -> evars, r - in - let evars, proof = - let proxy = + fun { state ; env ; + term1 = t ; ty1 = ty ; + cstr = (prop,cstr) ; evars } -> + let evars, rel = match cstr with + | None -> + let mkr = if prop then PropGlobal.mk_relation else TypeGlobal.mk_relation in + let evars, rty = mkr env evars ty in + new_cstr_evar evars env rty + | Some r -> evars, r + in + let evars, proof = + let proxy = if prop then PropGlobal.proper_proxy_type env else TypeGlobal.proper_proxy_type env - in - let evars, mty = app_poly_sort prop env evars proxy [| ty ; rel; t |] in - new_cstr_evar evars env mty - in - let res = Success { rew_car = ty; rew_from = t; rew_to = t; - rew_prf = RewPrf (rel, proof); rew_evars = evars } - in state, res + in + let evars, mty = app_poly_sort prop env evars proxy [| ty ; rel; t |] in + new_cstr_evar evars env mty + in + let res = Success { rew_car = ty; rew_from = t; rew_to = t; + rew_prf = RewPrf (rel, proof); rew_evars = evars } + in state, res } let progress (s : 'a pure_strategy) : 'a pure_strategy = { strategy = fun input -> - let state, res = s.strategy input in - match res with - | Fail -> state, Fail - | Identity -> state, Fail - | Success r -> state, Success r - } - + let state, res = s.strategy input in + match res with + | Fail -> state, Fail + | Identity -> state, Fail + | Success r -> state, Success r + } + let seq first snd : 'a pure_strategy = { strategy = fun ({ env ; unfresh ; cstr } as input) -> - let state, res = first.strategy input in - match res with - | Fail -> state, Fail - | Identity -> snd.strategy { input with state } - | Success res -> transitivity state env unfresh (fst cstr) res snd - } - + let state, res = first.strategy input in + match res with + | Fail -> state, Fail + | Identity -> snd.strategy { input with state } + | Success res -> transitivity state env unfresh (fst cstr) res snd + } + let choice fst snd : 'a pure_strategy = { strategy = fun input -> - let state, res = fst.strategy input in - match res with - | Fail -> snd.strategy { input with state } - | Identity | Success _ -> state, res - } + let state, res = fst.strategy input in + match res with + | Fail -> snd.strategy { input with state } + | Identity | Success _ -> state, res + } let try_ str : 'a pure_strategy = choice str id @@ -1357,7 +1357,7 @@ module Strategies = let fix (f : 'a pure_strategy -> 'a pure_strategy) : 'a pure_strategy = let rec aux input = (f { strategy = fun input -> check_interrupt aux input }).strategy input in { strategy = aux } - + let any (s : 'a pure_strategy) : 'a pure_strategy = fix (fun any -> try_ (seq s any)) @@ -1378,8 +1378,8 @@ module Strategies = let lemmas cs : 'a pure_strategy = List.fold_left (fun tac (l,l2r,by) -> - choice tac (apply_lemma l2r rewrite_unif_flags l by AllOccurrences)) - fail cs + choice tac (apply_lemma l2r rewrite_unif_flags l by AllOccurrences)) + fail cs let inj_open hint = (); fun sigma -> let ctx = UState.of_context_set hint.Autorewrite.rew_ctx in @@ -1388,51 +1388,51 @@ module Strategies = let old_hints (db : string) : 'a pure_strategy = let rules = Autorewrite.find_rewrites db in - lemmas - (List.map (fun hint -> (inj_open hint, hint.Autorewrite.rew_l2r, - hint.Autorewrite.rew_tac)) rules) + lemmas + (List.map (fun hint -> (inj_open hint, hint.Autorewrite.rew_l2r, + hint.Autorewrite.rew_tac)) rules) let hints (db : string) : 'a pure_strategy = { strategy = fun ({ term1 = t } as input) -> let t = EConstr.Unsafe.to_constr t in let rules = Autorewrite.find_matches db t in let lemma hint = (inj_open hint, hint.Autorewrite.rew_l2r, - hint.Autorewrite.rew_tac) in + hint.Autorewrite.rew_tac) in let lems = List.map lemma rules in (lemmas lems).strategy input - } + } let reduce (r : Redexpr.red_expr) : 'a pure_strategy = { strategy = - fun { state = state ; env = env ; term1 = t ; ty1 = ty ; cstr = cstr ; evars = evars } -> + fun { state = state ; env = env ; term1 = t ; ty1 = ty ; cstr = cstr ; evars = evars } -> let rfn, ckind = Redexpr.reduction_of_red_expr env r in let sigma = goalevars evars in - let (sigma, t') = rfn env sigma t in - if Termops.eq_constr sigma t' t then - state, Identity - else - state, Success { rew_car = ty; rew_from = t; rew_to = t'; - rew_prf = RewCast ckind; - rew_evars = sigma, cstrevars evars } - } - + let (sigma, t') = rfn env sigma t in + if Termops.eq_constr sigma t' t then + state, Identity + else + state, Success { rew_car = ty; rew_from = t; rew_to = t'; + rew_prf = RewCast ckind; + rew_evars = sigma, cstrevars evars } + } + let fold_glob c : 'a pure_strategy = { strategy = fun { state ; env ; term1 = t ; ty1 = ty ; cstr ; evars } -> (* let sigma, (c,_) = Tacinterp.interp_open_constr_with_bindings is env (goalevars evars) c in *) - let sigma, c = Pretyping.understand_tcc env (goalevars evars) c in - let unfolded = - try Tacred.try_red_product env sigma c - with e when CErrors.noncritical e -> + let sigma, c = Pretyping.understand_tcc env (goalevars evars) c in + let unfolded = + try Tacred.try_red_product env sigma c + with e when CErrors.noncritical e -> user_err Pp.(str "fold: the term is not unfoldable!") - in - try - let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in - let c' = Reductionops.nf_evar sigma c in - state, Success { rew_car = ty; rew_from = t; rew_to = c'; - rew_prf = RewCast DEFAULTcast; - rew_evars = (sigma, snd evars) } - with e when CErrors.noncritical e -> state, Fail - } - + in + try + let sigma = Unification.w_unify env sigma CONV ~flags:(Unification.elim_flags ()) unfolded t in + let c' = Reductionops.nf_evar sigma c in + state, Success { rew_car = ty; rew_from = t; rew_to = c'; + rew_prf = RewCast DEFAULTcast; + rew_evars = (sigma, snd evars) } + with e when CErrors.noncritical e -> state, Fail + } + end @@ -1450,19 +1450,19 @@ let rewrite_with l2r flags c occs : strategy = { strategy = unify_eqn rew l2r flags env (sigma, cstrs) None t in let app = apply_rule unify occs in - let strat = - Strategies.fix (fun aux -> - Strategies.choice app (subterm true default_flags aux)) + let strat = + Strategies.fix (fun aux -> + Strategies.choice app (subterm true default_flags aux)) in let _, res = strat.strategy { input with state = 0 } in ((), res) - } + } let apply_strategy (s : strategy) env unfresh concl (prop, cstr) evars = let ty = Retyping.get_type_of env (goalevars evars) concl in let _, res = s.strategy { state = () ; env ; unfresh ; - term1 = concl ; ty1 = ty ; - cstr = (prop, Some cstr) ; evars } in + term1 = concl ; ty1 = ty ; + cstr = (prop, Some cstr) ; evars } in res let solve_constraints env (evars,cstrs) = @@ -1483,14 +1483,14 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul let evdref = ref sigma in let evars = (!evdref, Evar.Set.empty) in let evars, cstr = - let prop, (evars, arrow) = + let prop, (evars, arrow) = if Sorts.is_prop sort then true, app_poly_sort true env evars impl [||] else false, app_poly_sort false env evars TypeGlobal.arrow [||] in match is_hyp with - | None -> - let evars, t = poly_inverse prop env evars (mkSort sort) arrow in - evars, (prop, t) + | None -> + let evars, t = poly_inverse prop env evars (mkSort sort) arrow in + evars, (prop, t) | Some _ -> evars, (prop, arrow) in let eq = apply_strategy strat env avoid concl cstr evars in @@ -1502,29 +1502,29 @@ let cl_rewrite_clause_aux ?(abs=None) strat env avoid sigma concl is_hyp : resul let evars' = solve_constraints env res.rew_evars in let newt = Reductionops.nf_evar evars' res.rew_to in let evars = (* Keep only original evars (potentially instantiated) and goal evars, - the rest has been defined and substituted already. *) - Evar.Set.fold - (fun ev acc -> - if not (Evd.is_defined acc ev) then - user_err ~hdr:"rewrite" - (str "Unsolved constraint remaining: " ++ spc () ++ + the rest has been defined and substituted already. *) + Evar.Set.fold + (fun ev acc -> + if not (Evd.is_defined acc ev) then + user_err ~hdr:"rewrite" + (str "Unsolved constraint remaining: " ++ spc () ++ Termops.pr_evar_info env acc (Evd.find acc ev)) - else Evd.remove acc ev) - cstrs evars' + else Evd.remove acc ev) + cstrs evars' in let res = match res.rew_prf with - | RewCast c -> None - | RewPrf (rel, p) -> - let p = nf_zeta env evars' (Reductionops.nf_evar evars' p) in - let term = - match abs with - | None -> p - | Some (t, ty) -> + | RewCast c -> None + | RewPrf (rel, p) -> + let p = nf_zeta env evars' (Reductionops.nf_evar evars' p) in + let term = + match abs with + | None -> p + | Some (t, ty) -> let t = Reductionops.nf_evar evars' t in let ty = Reductionops.nf_evar evars' ty in mkApp (mkLambda (make_annot (Name (Id.of_string "lemma")) Sorts.Relevant, ty, p), [| t |]) - in - let proof = match is_hyp with + in + let proof = match is_hyp with | None -> term | Some id -> mkApp (term, [| mkVar id |]) in Some proof @@ -1539,7 +1539,7 @@ let rec insert_dependent env sigma decl accu hyps = match hyps with else insert_dependent env sigma decl (ndecl :: accu) rem -let assert_replacing id newt tac = +let assert_replacing id newt tac = let prf = Proofview.Goal.enter begin fun gl -> let concl = Proofview.Goal.concl gl in let env = Proofview.Goal.env gl in @@ -1565,7 +1565,7 @@ let assert_replacing id newt tac = end in Proofview.tclTHEN prf (Proofview.tclFOCUS 2 2 tac) -let newfail n s = +let newfail n s = Proofview.tclZERO (Refiner.FailError (n, lazy s)) let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = @@ -1573,29 +1573,29 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = (* For compatibility *) let beta = Tactics.reduct_in_concl ~check:false (Reductionops.nf_betaiota, DEFAULTcast) in let beta_hyp id = Tactics.reduct_in_hyp ~check:false ~reorder:false Reductionops.nf_betaiota (id, InHyp) in - let treat sigma res = + let treat sigma res = match res with | None -> newfail 0 (str "Nothing to rewrite") | Some None -> if progress then newfail 0 (str"Failed to progress") - else Proofview.tclUNIT () + else Proofview.tclUNIT () | Some (Some res) -> let (undef, prf, newt) = res in let fold ev _ accu = if Evd.mem sigma ev then accu else ev :: accu in let gls = List.rev (Evd.fold_undefined fold undef []) in let gls = List.map Proofview.with_empty_state gls in match clause, prf with - | Some id, Some p -> + | Some id, Some p -> let tac = tclTHENLIST [ Refine.refine ~typecheck:true (fun h -> (h,p)); Proofview.Unsafe.tclNEWGOALS gls; ] in Proofview.Unsafe.tclEVARS undef <*> - tclTHENFIRST (assert_replacing id newt tac) (beta_hyp id) - | Some id, None -> + tclTHENFIRST (assert_replacing id newt tac) (beta_hyp id) + | Some id, None -> Proofview.Unsafe.tclEVARS undef <*> convert_hyp ~check:false ~reorder:false (LocalAssum (make_annot id Sorts.Relevant, newt)) <*> beta_hyp id - | None, Some p -> + | None, Some p -> Proofview.Unsafe.tclEVARS undef <*> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in @@ -1605,7 +1605,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = end in Refine.refine ~typecheck:true make <*> Proofview.Unsafe.tclNEWGOALS gls end - | None, None -> + | None, None -> Proofview.Unsafe.tclEVARS undef <*> convert_concl ~check:false newt DEFAULTcast in @@ -1639,7 +1639,7 @@ let cl_rewrite_clause_newtac ?abs ?origsigma ~progress strat clause = raise (RewriteFailure (Himsg.explain_pretype_error env evd e)) end -let tactic_init_setoid () = +let tactic_init_setoid () = try init_setoid (); Proofview.tclUNIT () with e when CErrors.noncritical e -> Tacticals.New.tclFAIL 0 (str"Setoid library not loaded") @@ -1650,9 +1650,9 @@ let cl_rewrite_clause_strat progress strat clause = (cl_rewrite_clause_newtac ~progress strat clause) (fun (e, info) -> match e with | RewriteFailure e -> - tclZEROMSG (str"setoid rewrite failed: " ++ e) - | Refiner.FailError (n, pp) -> - tclFAIL n (str"setoid rewrite failed: " ++ Lazy.force pp) + tclZEROMSG (str"setoid rewrite failed: " ++ e) + | Refiner.FailError (n, pp) -> + tclFAIL n (str"setoid rewrite failed: " ++ Lazy.force pp) | e -> Proofview.tclZERO ~info e)) (** Setoid rewriting when called with "setoid_rewrite" *) @@ -1663,7 +1663,7 @@ let cl_rewrite_clause l left2right occs clause = (** Setoid rewriting when called with "rewrite_strat" *) let cl_rewrite_clause_strat strat clause = cl_rewrite_clause_strat false strat clause - + let apply_glob_constr c l2r occs = (); fun ({ state = () ; env = env } as input) -> let c sigma = let (sigma, c) = Pretyping.understand_tcc env sigma c in @@ -1681,22 +1681,22 @@ let interp_glob_constr_list env = (* Syntax for rewriting with strategies *) -type unary_strategy = +type unary_strategy = Subterms | Subterm | Innermost | Outermost | Bottomup | Topdown | Progress | Try | Any | Repeat -type binary_strategy = +type binary_strategy = | Compose | Choice -type ('constr,'redexpr) strategy_ast = +type ('constr,'redexpr) strategy_ast = | StratId | StratFail | StratRefl | StratUnary of unary_strategy * ('constr,'redexpr) strategy_ast - | StratBinary of binary_strategy + | StratBinary of binary_strategy * ('constr,'redexpr) strategy_ast * ('constr,'redexpr) strategy_ast | StratConstr of 'constr * bool | StratTerms of 'constr list | StratHints of bool * string - | StratEval of 'redexpr + | StratEval of 'redexpr | StratFold of 'constr let rec map_strategy (f : 'a -> 'a2) (g : 'b -> 'b2) : ('a,'b) strategy_ast -> ('a2,'b2) strategy_ast = function @@ -1747,7 +1747,7 @@ let rec strategy_of_ast = function | StratId -> Strategies.id | StratFail -> Strategies.fail | StratRefl -> Strategies.refl - | StratUnary (f, s) -> + | StratUnary (f, s) -> let s' = strategy_of_ast s in let f' = match f with | Subterms -> all_subterms @@ -1774,12 +1774,12 @@ let rec strategy_of_ast = function (fun ({ state = () ; env } as input) -> let l' = interp_glob_constr_list env (List.map fst l) in (Strategies.lemmas l').strategy input) - } + } | StratEval r -> { strategy = (fun ({ state = () ; env ; evars } as input) -> let (sigma,r_interp) = Tacinterp.interp_redexp env (goalevars evars) r in (Strategies.reduce r_interp).strategy { input with - evars = (sigma,cstrevars evars) }) } + evars = (sigma,cstrevars evars) }) } | StratFold c -> Strategies.fold_glob (fst c) @@ -1862,7 +1862,7 @@ let proper_projection env sigma r ty = let mor, args = destApp sigma inst in let instarg = mkApp (r, rel_vect 0 (List.length ctx)) in let app = mkApp (PropGlobal.proper_proj env sigma, - Array.append args [| instarg |]) in + Array.append args [| instarg |]) in it_mkLambda_or_LetIn app ctx let declare_projection n instance_id r = @@ -1877,17 +1877,17 @@ let declare_projection n instance_id r = let typ = let n = let rec aux t = - match EConstr.kind sigma t with - | App (f, [| a ; a' ; rel; rel' |]) - when Termops.is_global sigma (PropGlobal.respectful_ref ()) f -> - succ (aux rel') - | _ -> 0 + match EConstr.kind sigma t with + | App (f, [| a ; a' ; rel; rel' |]) + when Termops.is_global sigma (PropGlobal.respectful_ref ()) f -> + succ (aux rel') + | _ -> 0 in let init = - match EConstr.kind sigma typ with - App (f, args) when Termops.is_global sigma (PropGlobal.respectful_ref ()) f -> - mkApp (f, fst (Array.chop (Array.length args - 2) args)) - | _ -> typ + match EConstr.kind sigma typ with + App (f, args) when Termops.is_global sigma (PropGlobal.respectful_ref ()) f -> + mkApp (f, fst (Array.chop (Array.length args - 2) args)) + | _ -> typ in aux init in let ctx,ccl = Reductionops.splay_prod_n env sigma (3 * n) typ @@ -1911,19 +1911,19 @@ let build_morphism_signature env sigma m = let rec aux t = match EConstr.kind sigma t with | Prod (na, a, b) -> - None :: aux b - | _ -> [] + None :: aux b + | _ -> [] in aux t in - let evars, t', sig_, cstrs = + let evars, t', sig_, cstrs = PropGlobal.build_signature (sigma, Evar.Set.empty) env t cstrs None in let evd = ref evars in let _ = List.iter (fun (ty, rel) -> Option.iter (fun rel -> - let default = e_app_poly env evd PropGlobal.default_relation [| ty; rel |] in - ignore(e_new_cstr_evar env evd default)) - rel) + let default = e_app_poly env evd PropGlobal.default_relation [| ty; rel |] in + ignore(e_new_cstr_evar env evd default)) + rel) cstrs in let morph = e_app_poly env evd (PropGlobal.proper_type env) [| t; sig_; m |] in @@ -2062,14 +2062,14 @@ let unification_rewrite l2r c1 c2 sigma prf car rel but env = (* ~flags:(false,true) to allow to mark occurrences that must not be rewritten simply by replacing them with let-defined definitions in the context *) - Unification.w_unify_to_subterm + Unification.w_unify_to_subterm ~flags:rewrite_unif_flags env sigma ((if l2r then c1 else c2),but) with | ex when Pretype_errors.precatchable_exception ex -> - (* ~flags:(true,true) to make Ring work (since it really + (* ~flags:(true,true) to make Ring work (since it really exploits conversion) *) - Unification.w_unify_to_subterm + Unification.w_unify_to_subterm ~flags:rewrite_conv_unif_flags env sigma ((if l2r then c1 else c2),but) in @@ -2112,15 +2112,15 @@ let general_s_rewrite cl l2r occs (c,l) ~new_goals = let strat = { strategy = fun ({ state = () } as input) -> let _, res = substrat.strategy { input with state = 0 } in (), res - } + } in let origsigma = Tacmach.New.project gl in tactic_init_setoid () <*> Proofview.tclOR (tclPROGRESS - (tclTHEN + (tclTHEN (Proofview.Unsafe.tclEVARS evd) - (cl_rewrite_clause_newtac ~progress:true ~abs:(Some abs) ~origsigma strat cl))) + (cl_rewrite_clause_newtac ~progress:true ~abs:(Some abs) ~origsigma strat cl))) (fun (e, info) -> match e with | RewriteFailure e -> tclFAIL 0 (str"setoid rewrite failed: " ++ e) @@ -2147,7 +2147,7 @@ let setoid_proof ty fn fallback = let rel, _, _ = decompose_app_rel env sigma concl in let (sigma, t) = Typing.type_of env sigma rel in let car = snd (List.hd (fst (Reductionops.splay_prod env sigma t))) in - (try init_relation_classes () with _ -> raise Not_found); + (try init_relation_classes () with _ -> raise Not_found); fn env sigma car rel with e -> Proofview.tclZERO e end @@ -2157,18 +2157,18 @@ let setoid_proof ty fn fallback = fallback begin function (e', info) -> match e' with | Hipattern.NoEquationFound -> - begin match e with - | (Not_found, _) -> - let rel, _, _ = decompose_app_rel env sigma concl in - not_declared env sigma ty rel - | (e, info) -> Proofview.tclZERO ~info e + begin match e with + | (Not_found, _) -> + let rel, _, _ = decompose_app_rel env sigma concl in + not_declared env sigma ty rel + | (e, info) -> Proofview.tclZERO ~info e end | e' -> Proofview.tclZERO ~info e' end end end -let tac_open ((evm,_), c) tac = +let tac_open ((evm,_), c) tac = (tclTHEN (Proofview.Unsafe.tclEVARS evm) (tac c)) let poly_proof getp gett env evm car rel = @@ -2178,32 +2178,32 @@ let poly_proof getp gett env evm car rel = let setoid_reflexivity = setoid_proof "reflexive" - (fun env evm car rel -> + (fun env evm car rel -> tac_open (poly_proof PropGlobal.get_reflexive_proof - TypeGlobal.get_reflexive_proof - env evm car rel) - (fun c -> tclCOMPLETE (apply c))) + TypeGlobal.get_reflexive_proof + env evm car rel) + (fun c -> tclCOMPLETE (apply c))) (reflexivity_red true) let setoid_symmetry = setoid_proof "symmetric" - (fun env evm car rel -> + (fun env evm car rel -> tac_open - (poly_proof PropGlobal.get_symmetric_proof TypeGlobal.get_symmetric_proof - env evm car rel) - (fun c -> apply c)) + (poly_proof PropGlobal.get_symmetric_proof TypeGlobal.get_symmetric_proof + env evm car rel) + (fun c -> apply c)) (symmetry_red true) - + let setoid_transitivity c = setoid_proof "transitive" (fun env evm car rel -> tac_open (poly_proof PropGlobal.get_transitive_proof TypeGlobal.get_transitive_proof - env evm car rel) - (fun proof -> match c with - | None -> eapply proof - | Some c -> apply_with_bindings (proof,ImplicitBindings [ c ]))) + env evm car rel) + (fun proof -> match c with + | None -> eapply proof + | Some c -> apply_with_bindings (proof,ImplicitBindings [ c ]))) (transitivity_red true c) - + let setoid_symmetry_in id = let open Tacmach.New in Proofview.Goal.enter begin fun gl -> @@ -2230,16 +2230,16 @@ let _ = Hook.set Tactics.setoid_symmetry setoid_symmetry let _ = Hook.set Tactics.setoid_symmetry_in setoid_symmetry_in let _ = Hook.set Tactics.setoid_transitivity setoid_transitivity -let get_lemma_proof f env evm x y = +let get_lemma_proof f env evm x y = let (evm, _), c = f env (evm,Evar.Set.empty) x y in evm, c let get_reflexive_proof = get_lemma_proof PropGlobal.get_reflexive_proof -let get_symmetric_proof = +let get_symmetric_proof = get_lemma_proof PropGlobal.get_symmetric_proof -let get_transitive_proof = +let get_transitive_proof = get_lemma_proof PropGlobal.get_transitive_proof - + diff --git a/plugins/ltac/rewrite.mli b/plugins/ltac/rewrite.mli index 8e0b0a8003..576ed686d4 100644 --- a/plugins/ltac/rewrite.mli +++ b/plugins/ltac/rewrite.mli @@ -23,25 +23,25 @@ open Tacinterp type rewrite_attributes val rewrite_attributes : rewrite_attributes Attributes.attribute -type unary_strategy = +type unary_strategy = Subterms | Subterm | Innermost | Outermost | Bottomup | Topdown | Progress | Try | Any | Repeat -type binary_strategy = +type binary_strategy = | Compose | Choice -type ('constr,'redexpr) strategy_ast = +type ('constr,'redexpr) strategy_ast = | StratId | StratFail | StratRefl | StratUnary of unary_strategy * ('constr,'redexpr) strategy_ast - | StratBinary of binary_strategy + | StratBinary of binary_strategy * ('constr,'redexpr) strategy_ast * ('constr,'redexpr) strategy_ast | StratConstr of 'constr * bool | StratTerms of 'constr list | StratHints of bool * string - | StratEval of 'redexpr + | StratEval of 'redexpr | StratFold of 'constr -type rewrite_proof = +type rewrite_proof = | RewPrf of constr * constr | RewCast of Constr.cast_kind diff --git a/plugins/ltac/taccoerce.ml b/plugins/ltac/taccoerce.ml index da89a027e2..a57cc76faa 100644 --- a/plugins/ltac/taccoerce.ml +++ b/plugins/ltac/taccoerce.ml @@ -191,7 +191,7 @@ let id_of_name = function | None -> fail () | Some c -> match EConstr.kind sigma c with - | Var id -> id + | Var id -> id | Meta m -> id_of_name (Evd.meta_name sigma m) | Evar (kn,_) -> begin match Evd.evar_ident kn sigma with @@ -201,12 +201,12 @@ let id_of_name = function | Const (cst,_) -> Label.to_id (Constant.label cst) | Construct (cstr,_) -> let ref = GlobRef.ConstructRef cstr in - let basename = Nametab.basename_of_global ref in - basename + let basename = Nametab.basename_of_global ref in + basename | Ind (ind,_) -> let ref = GlobRef.IndRef ind in - let basename = Nametab.basename_of_global ref in - basename + let basename = Nametab.basename_of_global ref in + basename | Sort s -> begin match ESorts.kind sigma s with diff --git a/plugins/ltac/tacenv.ml b/plugins/ltac/tacenv.ml index 6a5ab55604..8bafbb7ea3 100644 --- a/plugins/ltac/tacenv.ml +++ b/plugins/ltac/tacenv.ml @@ -107,7 +107,7 @@ let interp_ml_tactic { mltac_name = s; mltac_index = i } = let () = if Array.length tacs <= i then raise Not_found in tacs.(i) with Not_found -> - CErrors.user_err + CErrors.user_err (str "The tactic " ++ pr_tacname s ++ str " is not installed.") (***************************************************************************) diff --git a/plugins/ltac/tacintern.ml b/plugins/ltac/tacintern.ml index 63559cf488..4dc2ade7a1 100644 --- a/plugins/ltac/tacintern.ml +++ b/plugins/ltac/tacintern.ml @@ -279,11 +279,11 @@ let intern_destruction_arg ist = function | clear,ElimOnAnonHyp n as x -> x | clear,ElimOnIdent {loc;v=id} -> if !strict_check then - (* If in a defined tactic, no intros-until *) + (* If in a defined tactic, no intros-until *) let c, p = intern_constr ist (make @@ CRef (qualid_of_ident id, None)) in - match DAst.get c with + match DAst.get c with | GVar id -> clear,ElimOnIdent (make ?loc:c.loc id) - | _ -> clear,ElimOnConstr ((c, p), NoBindings) + | _ -> clear,ElimOnConstr ((c, p), NoBindings) else clear,ElimOnIdent (make ?loc id) @@ -401,13 +401,13 @@ let dump_glob_red_expr = function | Unfold occs -> List.iter (fun (_, r) -> try Dumpglob.add_glob ?loc:r.loc - (Smartlocate.smart_global r) + (Smartlocate.smart_global r) with e when CErrors.noncritical e -> ()) occs | Cbv grf | Lazy grf -> List.iter (fun r -> try Dumpglob.add_glob ?loc:r.loc - (Smartlocate.smart_global r) + (Smartlocate.smart_global r) with e when CErrors.noncritical e -> ()) grf.rConst | _ -> () @@ -525,18 +525,18 @@ let rec intern_atomic lf ist x = intern_constr_gen false (not (Option.is_empty otac)) ist c) | TacGeneralize cl -> TacGeneralize (List.map (fun (c,na) -> - intern_constr_with_occurrences ist c, + intern_constr_with_occurrences ist c, intern_name lf ist na) cl) | TacLetTac (ev,na,c,cls,b,eqpat) -> let na = intern_name lf ist na in TacLetTac (ev,na,intern_constr ist c, (clause_app (intern_hyp_location ist) cls),b, - (Option.map (intern_intro_pattern_naming_loc lf ist) eqpat)) + (Option.map (intern_intro_pattern_naming_loc lf ist) eqpat)) (* Derived basic tactics *) | TacInductionDestruct (ev,isrec,(l,el)) -> TacInductionDestruct (ev,isrec,(List.map (fun (c,(ipato,ipats),cls) -> - (intern_destruction_arg ist c, + (intern_destruction_arg ist c, (Option.map (intern_intro_pattern_naming_loc lf ist) ipato, Option.map (intern_or_and_intro_pattern_loc lf ist) ipats), Option.map (clause_app (intern_hyp_location ist)) cls)) l, @@ -557,7 +557,7 @@ let rec intern_atomic lf ist x = TacChange (check,None, (if is_onhyps && is_onconcl then intern_type ist c else intern_constr ist c), - clause_app (intern_hyp_location ist) cl) + clause_app (intern_hyp_location ist) cl) | TacChange (check,Some p,c,cl) -> let { ltacvars } = ist in let metas,pat = intern_typed_pattern ist ~as_type:false ~ltacvars p in @@ -565,15 +565,15 @@ let rec intern_atomic lf ist x = let ltacvars = List.fold_left fold ltacvars metas in let ist' = { ist with ltacvars } in TacChange (check,Some pat,intern_constr ist' c, - clause_app (intern_hyp_location ist) cl) + clause_app (intern_hyp_location ist) cl) (* Equality and inversion *) | TacRewrite (ev,l,cl,by) -> TacRewrite - (ev, - List.map (fun (b,m,c) -> (b,m,intern_constr_with_bindings_arg ist c)) l, - clause_app (intern_hyp_location ist) cl, - Option.map (intern_pure_tactic ist) by) + (ev, + List.map (fun (b,m,c) -> (b,m,intern_constr_with_bindings_arg ist c)) l, + clause_app (intern_hyp_location ist) cl, + Option.map (intern_pure_tactic ist) by) | TacInversion (inv,hyp) -> TacInversion (intern_inversion_strength lf ist inv, intern_quantified_hypothesis ist hyp) @@ -590,7 +590,7 @@ and intern_tactic_seq onlytac ist = function let ltacvars = Id.Set.union (extract_let_names l) ist.ltacvars in let ist' = { ist with ltacvars } in let l = List.map (fun (n,b) -> - (n,intern_tacarg !strict_check false (if isrec then ist' else ist) b)) l in + (n,intern_tacarg !strict_check false (if isrec then ist' else ist) b)) l in ist.ltacvars, TacLetIn (isrec,l,intern_tactic onlytac ist' u) | TacMatchGoal (lz,lr,lmr) -> @@ -615,13 +615,13 @@ and intern_tactic_seq onlytac ist = function ist.ltacvars , TacExtendTac (Array.map (intern_pure_tactic ist) tf, intern_pure_tactic ist t, - Array.map (intern_pure_tactic ist) tl) + Array.map (intern_pure_tactic ist) tl) | TacThens3parts (t1,tf,t2,tl) -> let lfun', t1 = intern_tactic_seq onlytac ist t1 in let ist' = { ist with ltacvars = lfun' } in (* Que faire en cas de (tac complexe avec Match et Thens; tac2) ?? *) lfun', TacThens3parts (t1,Array.map (intern_pure_tactic ist') tf,intern_pure_tactic ist' t2, - Array.map (intern_pure_tactic ist') tl) + Array.map (intern_pure_tactic ist') tl) | TacThens (t,tl) -> let lfun', t = intern_tactic_seq true ist t in let ist' = { ist with ltacvars = lfun' } in diff --git a/plugins/ltac/tacinterp.ml b/plugins/ltac/tacinterp.ml index c252372f21..9633c9bd77 100644 --- a/plugins/ltac/tacinterp.ml +++ b/plugins/ltac/tacinterp.ml @@ -240,7 +240,7 @@ let append_trace trace v = (* Dynamically check that an argument is a tactic *) let coerce_to_tactic loc id v = - let fail () = user_err ?loc + let fail () = user_err ?loc (str "Variable " ++ Id.print id ++ str " should be bound to a tactic.") in if has_type v (topwit wit_tacvalue) then @@ -472,8 +472,8 @@ let interp_fresh_id ist env sigma l = if List.is_empty l then default_fresh_id else let s = - String.concat "" (List.map (function - | ArgArg s -> s + String.concat "" (List.map (function + | ArgArg s -> s | ArgVar {v=id} -> Id.to_string (extract_ident ist env sigma id)) l) in let s = if CLexer.is_keyword s then s^"0" else s in Id.of_string s in @@ -694,7 +694,7 @@ let interp_red_expr ist env sigma = function sigma , Pattern l_interp | Simpl (f,o) -> sigma , Simpl (interp_flag ist env sigma f, - Option.map (interp_closed_typed_pattern_with_occurrences ist env sigma) o) + Option.map (interp_closed_typed_pattern_with_occurrences ist env sigma) o) | CbvVm o -> sigma , CbvVm (Option.map (interp_closed_typed_pattern_with_occurrences ist env sigma) o) | CbvNative o -> @@ -709,23 +709,23 @@ let interp_may_eval f ist env sigma = function redfun env sigma c_interp | ConstrContext ({loc;v=s},c) -> (try - let (sigma,ic) = f ist env sigma c in + let (sigma,ic) = f ist env sigma c in let ctxt = try_interp_ltac_var coerce_to_constr_context ist (Some (env, sigma)) (make ?loc s) in - let ctxt = EConstr.Unsafe.to_constr ctxt in + let ctxt = EConstr.Unsafe.to_constr ctxt in let ic = EConstr.Unsafe.to_constr ic in - let c = subst_meta [Constr_matching.special_meta,ic] ctxt in + let c = subst_meta [Constr_matching.special_meta,ic] ctxt in Typing.solve_evars env sigma (EConstr.of_constr c) with - | Not_found -> - user_err ?loc ~hdr:"interp_may_eval" - (str "Unbound context identifier" ++ Id.print s ++ str".")) + | Not_found -> + user_err ?loc ~hdr:"interp_may_eval" + (str "Unbound context identifier" ++ Id.print s ++ str".")) | ConstrTypeOf c -> let (sigma,c_interp) = f ist env sigma c in let (sigma, t) = Typing.type_of ~refresh:true env sigma c_interp in (sigma, t) | ConstrTerm c -> try - f ist env sigma c + f ist env sigma c with reraise -> let reraise = CErrors.push reraise in (* spiwack: to avoid unnecessary modifications of tacinterp, as this @@ -909,7 +909,7 @@ let interp_destruction_arg ist gl arg = end | keep,ElimOnAnonHyp n as x -> x | keep,ElimOnIdent {loc;v=id} -> - let error () = user_err ?loc + let error () = user_err ?loc (strbrk "Cannot coerce " ++ Id.print id ++ strbrk " neither to a quantified hypothesis nor to a term.") in @@ -941,10 +941,10 @@ let interp_destruction_arg ist gl arg = | None -> error () | Some c -> keep,ElimOnConstr (fun env sigma -> (sigma, (c,NoBindings))) with Not_found -> - (* We were in non strict (interactive) mode *) - if Tactics.is_quantified_hypothesis id gl then + (* We were in non strict (interactive) mode *) + if Tactics.is_quantified_hypothesis id gl then keep,ElimOnIdent (make ?loc id) - else + else let c = (DAst.make ?loc @@ GVar id,Some (make @@ CRef (qualid_of_ident ?loc id,None))) in let f env sigma = let (sigma,c) = interp_open_constr ist env sigma c in @@ -995,11 +995,11 @@ let rec read_match_goal_hyps lfun ist env sigma lidh = function | (Hyp ({loc;v=na} as locna,mp))::tl -> let lidh' = Name.fold_right cons_and_check_name na lidh in Hyp (locna,read_pattern lfun ist env sigma mp):: - (read_match_goal_hyps lfun ist env sigma lidh' tl) + (read_match_goal_hyps lfun ist env sigma lidh' tl) | (Def ({loc;v=na} as locna,mv,mp))::tl -> let lidh' = Name.fold_right cons_and_check_name na lidh in Def (locna,read_pattern lfun ist env sigma mv, read_pattern lfun ist env sigma mp):: - (read_match_goal_hyps lfun ist env sigma lidh' tl) + (read_match_goal_hyps lfun ist env sigma lidh' tl) | [] -> [] (* Reads the rules of a Match Context or a Match *) @@ -1060,7 +1060,7 @@ let rec val_interp ist ?(appl=UnnamedAppl) (tac:glob_tactic_expr) : Val.t Ftacti let ist = { ist with extra = TacStore.set ist.extra f_debug v } in value_interp ist >>= fun v -> return (name_vfun appl v) in - Tactic_debug.debug_prompt lev tac eval + Tactic_debug.debug_prompt lev tac eval | _ -> value_interp ist >>= fun v -> return (name_vfun appl v) @@ -1117,7 +1117,7 @@ and eval_tactic ist tac : unit Proofview.tactic = match tac with | TacThens (t1,tl) -> Tacticals.New.tclTHENS (interp_tactic ist t1) (List.map (interp_tactic ist) tl) | TacThens3parts (t1,tf,t,tl) -> Tacticals.New.tclTHENS3PARTS (interp_tactic ist t1) - (Array.map (interp_tactic ist) tf) (interp_tactic ist t) (Array.map (interp_tactic ist) tl) + (Array.map (interp_tactic ist) tf) (interp_tactic ist t) (Array.map (interp_tactic ist) tl) | TacDo (n,tac) -> Tacticals.New.tclDO (interp_int_or_var ist n) (interp_tactic ist tac) | TacTimeout (n,tac) -> Tacticals.New.tclTIMEOUT (interp_int_or_var ist n) (interp_tactic ist tac) | TacTime (s,tac) -> Tacticals.New.tclTIME s (interp_tactic ist tac) @@ -1276,9 +1276,9 @@ and interp_app loc ist fv largs : Val.t Ftactic.t = | (VFun(appl,trace,olfun,(_::_ as var),body) |VFun(appl,trace,olfun,([] as var), (TacFun _|TacLetIn _|TacMatchGoal _|TacMatch _| TacArg _ as body))) -> - let (extfun,lvar,lval)=head_with_value (var,largs) in + let (extfun,lvar,lval)=head_with_value (var,largs) in let fold accu (id, v) = Id.Map.add id v accu in - let newlfun = List.fold_left fold olfun extfun in + let newlfun = List.fold_left fold olfun extfun in if List.is_empty lvar then begin wrap_error begin @@ -1291,9 +1291,9 @@ and interp_app loc ist fv largs : Val.t Ftactic.t = (catch_error_tac trace (val_interp ist body)) >>= fun v -> Ftactic.return (name_vfun (push_appl appl largs) v) end - begin fun (e, info) -> + begin fun (e, info) -> Proofview.tclLIFT (debugging_exception_step ist false e (fun () -> str "evaluation")) <*> - Proofview.tclZERO ~info e + Proofview.tclZERO ~info e end end >>= fun v -> (* No errors happened, we propagate the trace *) @@ -1604,10 +1604,10 @@ and interp_atomic ist tac : unit Proofview.tactic = Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in let sigma = project gl in - let l = List.map (fun (k,c) -> + let l = List.map (fun (k,c) -> let loc, f = interp_open_constr_with_bindings_loc ist c in (k,(make ?loc f))) cb - in + in let sigma,tac = match cl with | None -> sigma, Tactics.apply_with_delayed_bindings_gen a ev l | Some cl -> @@ -1619,7 +1619,7 @@ and interp_atomic ist tac : unit Proofview.tactic = | TacElim (ev,(keep,cb),cbo) -> Proofview.Goal.enter begin fun gl -> let env = Proofview.Goal.env gl in - let sigma = project gl in + let sigma = project gl in let sigma, cb = interp_open_constr_with_bindings ist env sigma cb in let sigma, cbo = Option.fold_left_map (interp_open_constr_with_bindings ist env) sigma cbo in let named_tac = @@ -1646,7 +1646,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let env = pf_env gl in let f sigma (id,n,c) = let (sigma,c_interp) = interp_type ist env sigma c in - sigma , (interp_ident ist env sigma id,n,c_interp) in + sigma , (interp_ident ist env sigma id,n,c_interp) in let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> f sigma c) l (project gl) in @@ -1660,8 +1660,8 @@ and interp_atomic ist tac : unit Proofview.tactic = Proofview.Goal.enter begin fun gl -> let env = pf_env gl in let f sigma (id,c) = - let (sigma,c_interp) = interp_type ist env sigma c in - sigma , (interp_ident ist env sigma id,c_interp) in + let (sigma,c_interp) = interp_type ist env sigma c in + sigma , (interp_ident ist env sigma id,c_interp) in let (sigma,l_interp) = Evd.MonadR.List.map_right (fun c sigma -> f sigma c) l (project gl) in @@ -1728,7 +1728,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let (sigma',c) = interp_pure_open_constr ist env sigma c in name_atomic ~env (TacLetTac(ev,na,c,clp,b,eqpat)) - (Tacticals.New.tclWITHHOLES ev + (Tacticals.New.tclWITHHOLES ev (let_pat_tac b (interp_name ist env sigma na) (sigma,c) clp eqpat) sigma') end @@ -1782,11 +1782,11 @@ and interp_atomic ist tac : unit Proofview.tactic = | _ -> false in let c_interp patvars env sigma = - let lfun' = Id.Map.fold (fun id c lfun -> - Id.Map.add id (Value.of_constr c) lfun) - patvars ist.lfun - in - let ist = { ist with lfun = lfun' } in + let lfun' = Id.Map.fold (fun id c lfun -> + Id.Map.add id (Value.of_constr c) lfun) + patvars ist.lfun + in + let ist = { ist with lfun = lfun' } in if is_onhyps && is_onconcl then interp_type ist env sigma c else interp_constr ist env sigma c @@ -1804,7 +1804,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let to_catch = function Not_found -> true | e -> CErrors.is_anomaly e in let c_interp patvars env sigma = let lfun' = Id.Map.fold (fun id c lfun -> - Id.Map.add id (Value.of_constr c) lfun) + Id.Map.add id (Value.of_constr c) lfun) patvars ist.lfun in let env = ensure_freshness env in @@ -1826,7 +1826,7 @@ and interp_atomic ist tac : unit Proofview.tactic = let f env sigma = interp_open_constr_with_bindings ist env sigma c in - (b,m,keep,f)) l in + (b,m,keep,f)) l in let env = Proofview.Goal.env gl in let sigma = project gl in let cl = interp_clause ist env sigma cl in diff --git a/plugins/ltac/tacsubst.ml b/plugins/ltac/tacsubst.ml index bf5d49f678..e864d31da4 100644 --- a/plugins/ltac/tacsubst.ml +++ b/plugins/ltac/tacsubst.ml @@ -161,9 +161,9 @@ let rec subst_atomic subst (t:glob_atomic_tactic_expr) = match t with (* Equality and inversion *) | TacRewrite (ev,l,cl,by) -> TacRewrite (ev, - List.map (fun (b,m,c) -> - b,m,subst_glob_with_bindings_arg subst c) l, - cl,Option.map (subst_tactic subst) by) + List.map (fun (b,m,c) -> + b,m,subst_glob_with_bindings_arg subst c) l, + cl,Option.map (subst_tactic subst) by) | TacInversion (DepInversion (k,c,l),hyp) -> TacInversion (DepInversion (k,Option.map (subst_glob_constr subst) c,l),hyp) | TacInversion (NonDepInversion _,_) as x -> x @@ -189,13 +189,13 @@ and subst_tactic subst (t:glob_tactic_expr) = match t with | TacDispatch tl -> TacDispatch (List.map (subst_tactic subst) tl) | TacExtendTac (tf,t,tl) -> TacExtendTac (Array.map (subst_tactic subst) tf, - subst_tactic subst t, + subst_tactic subst t, Array.map (subst_tactic subst) tl) | TacThens (t,tl) -> TacThens (subst_tactic subst t, List.map (subst_tactic subst) tl) | TacThens3parts (t1,tf,t2,tl) -> TacThens3parts (subst_tactic subst t1,Array.map (subst_tactic subst) tf, - subst_tactic subst t2,Array.map (subst_tactic subst) tl) + subst_tactic subst t2,Array.map (subst_tactic subst) tl) | TacDo (n,tac) -> TacDo (n,subst_tactic subst tac) | TacTimeout (n,tac) -> TacTimeout (n,subst_tactic subst tac) | TacTime (s,tac) -> TacTime (s,subst_tactic subst tac) diff --git a/plugins/ltac/tactic_matching.ml b/plugins/ltac/tactic_matching.ml index d008f9da1f..eabfe2f540 100644 --- a/plugins/ltac/tactic_matching.ml +++ b/plugins/ltac/tactic_matching.ml @@ -85,7 +85,7 @@ let is_empty_subst (ln,lm) = would ensure consistency. *) let equal_instances env sigma (ctx',c') (ctx,c) = (* How to compare instances? Do we want the terms to be convertible? - unifiable? Do we want the universe levels to be relevant? + unifiable? Do we want the universe levels to be relevant? (historically, conv_x is used) *) CList.equal Id.equal ctx ctx' && Reductionops.is_conv env sigma c' c @@ -230,11 +230,11 @@ module PatternMatching (E:StaticEnvironment) = struct (** [pattern_match_term refresh pat term lhs] returns the possible matchings of [term] with the pattern [pat => lhs]. If refresh is true, refreshes the universes of [term]. *) - let pattern_match_term refresh pat term lhs = + let pattern_match_term refresh pat term lhs = (* let term = if refresh then Termops.refresh_universes_strict term else term in *) match pat with | Term p -> - begin + begin try put_subst (Constr_matching.extended_matches E.env E.sigma p term) <*> return lhs diff --git a/plugins/ltac/tactic_option.ml b/plugins/ltac/tactic_option.ml index 21e02d4c04..da57f51ca3 100644 --- a/plugins/ltac/tactic_option.ml +++ b/plugins/ltac/tactic_option.ml @@ -34,19 +34,19 @@ let declare_tactic_option ?(default=Tacexpr.TacId []) name = let input : bool * Tacexpr.glob_tactic_expr -> obj = declare_object { (default_object name) with - cache_function = cache; - load_function = (fun _ -> load); - open_function = (fun _ -> load); - classify_function = (fun (local, tac) -> - if local then Dispose else Substitute (local, tac)); - subst_function = subst} + cache_function = cache; + load_function = (fun _ -> load); + open_function = (fun _ -> load); + classify_function = (fun (local, tac) -> + if local then Dispose else Substitute (local, tac)); + subst_function = subst} in let put local tac = set_default_tactic local tac; Lib.add_anonymous_leaf (input (local, tac)) in let get () = !locality, Tacinterp.eval_tactic !default_tactic in - let print () = + let print () = Pptactic.pr_glob_tactic (Global.env ()) !default_tactic_expr ++ (if !locality then str" (locally defined)" else str" (globally defined)") in diff --git a/plugins/ltac/tactic_option.mli b/plugins/ltac/tactic_option.mli index 637dd238fe..9705d225d4 100644 --- a/plugins/ltac/tactic_option.mli +++ b/plugins/ltac/tactic_option.mli @@ -11,7 +11,7 @@ open Tacexpr open Vernacexpr -val declare_tactic_option : ?default:Tacexpr.glob_tactic_expr -> string -> +val declare_tactic_option : ?default:Tacexpr.glob_tactic_expr -> string -> (* put *) (locality_flag -> glob_tactic_expr -> unit) * (* get *) (unit -> locality_flag * unit Proofview.tactic) * (* print *) (unit -> Pp.t) diff --git a/plugins/micromega/certificate.ml b/plugins/micromega/certificate.ml index 24039c93c6..82c2be582b 100644 --- a/plugins/micromega/certificate.ml +++ b/plugins/micromega/certificate.ml @@ -96,7 +96,7 @@ let rec fixpoint f x = if (=) y' x then y' else fixpoint f y' -let rec_simpl_cone n_spec e = +let rec_simpl_cone n_spec e = let simpl_cone = Mc.simpl_cone n_spec.zero n_spec.unit n_spec.mult n_spec.eqb in @@ -107,21 +107,21 @@ let rec_simpl_cone n_spec e = simpl_cone (Mc.PsatzAdd (rec_simpl_cone t1, rec_simpl_cone t2)) | x -> simpl_cone x in rec_simpl_cone e - - + + let simplify_cone n_spec c = fixpoint (rec_simpl_cone n_spec) c -(* The binding with Fourier might be a bit obsolete +(* The binding with Fourier might be a bit obsolete -- how does it handle equalities ? *) (* Certificates are elements of the cone such that P = 0 *) (* To begin with, we search for certificates of the form: - a1.p1 + ... an.pn + b1.q1 +... + bn.qn + c = 0 + a1.p1 + ... an.pn + b1.q1 +... + bn.qn + c = 0 where pi >= 0 qi > 0 - ai >= 0 + ai >= 0 bi >= 0 Sum bi + c >= 1 This is a linear problem: each monomial is considered as a variable. @@ -135,7 +135,7 @@ let simplify_cone n_spec c = fixpoint (rec_simpl_cone n_spec) c let constrain_variable v l = let coeffs = List.fold_left (fun acc p -> (Vect.get v p.coeffs)::acc) [] l in { coeffs = Vect.from_list ((Big_int zero_big_int):: (Big_int zero_big_int):: (List.rev coeffs)) ; - op = Eq ; + op = Eq ; cst = Big_int zero_big_int } @@ -143,10 +143,10 @@ let constrain_variable v l = let constrain_constant l = let coeffs = List.fold_left (fun acc p -> minus_num p.cst ::acc) [] l in { coeffs = Vect.from_list ((Big_int zero_big_int):: (Big_int unit_big_int):: (List.rev coeffs)) ; - op = Eq ; + op = Eq ; cst = Big_int zero_big_int } -let positivity l = +let positivity l = let rec xpositivity i l = match l with | [] -> [] @@ -169,7 +169,7 @@ let cstr_of_poly (p,o) = let variables_of_cstr c = Vect.variables c.coeffs -(* If the certificate includes at least one strict inequality, +(* If the certificate includes at least one strict inequality, the obtained polynomial can also be 0 *) let build_dual_linear_system l = @@ -486,7 +486,7 @@ let square_of_var i = let x = LinPoly.var i in ((LinPoly.product x x,Ge),(ProofFormat.Square x)) - + (** [nlinear_preprocess sys] augments the system [sys] by performing some limited non-linear reasoning. For instance, it asserts that the x² ≥0 but also that if c₁ ≥ 0 ∈ sys and c₂ ≥ 0 ∈ sys then c₁ × c₂ ≥ 0. The resulting system is linearised. @@ -510,7 +510,7 @@ let nlinear_preprocess (sys:WithProof.t list) = let sys = ISet.fold (fun i acc -> square_of_var i :: acc) collect_vars sys in let sys = sys @ (all_pairs WithProof.product sys) in - + if debug then begin Printf.fprintf stdout "Preprocessed\n"; List.iter (fun s -> Printf.fprintf stdout "%a\n" WithProof.output s) sys; @@ -545,12 +545,12 @@ let linear_prover_with_cert prfdepth sys = | Some cert -> Prf (ProofFormat.cmpl_prf_rule Mc.normQ CamlToCoq.q (List.mapi (fun i e -> i) sys) cert) -(* The prover is (probably) incomplete -- +(* The prover is (probably) incomplete -- only searching for naive cutting planes *) open Sos_types -let rec scale_term t = +let rec scale_term t = match t with | Zero -> unit_big_int , Zero | Const n -> (denominator n) , Const (Big_int (numerator n)) @@ -564,7 +564,7 @@ let rec scale_term t = if Int.equal (compare_big_int e unit_big_int) 0 then (unit_big_int, Add (y1,y2)) else e, Add (Mul(Const (Big_int s2'), y1), - Mul (Const (Big_int s1'), y2)) + Mul (Const (Big_int s1'), y2)) | Sub _ -> failwith "scale term: not implemented" | Mul(y,z) -> let s1,y1 = scale_term y and s2,y2 = scale_term z in mult_big_int s1 s2 , Mul (y1, y2) @@ -615,14 +615,14 @@ let rec term_to_q_expr = function let term_to_q_pol e = Mc.norm_aux (Ml2C.q (Int 0)) (Ml2C.q (Int 1)) Mc.qplus Mc.qmult Mc.qminus Mc.qopp Mc.qeq_bool (term_to_q_expr e) -let rec product l = +let rec product l = match l with | [] -> Mc.PsatzZ | [i] -> Mc.PsatzIn (Ml2C.nat i) | i ::l -> Mc.PsatzMulE(Mc.PsatzIn (Ml2C.nat i), product l) -let q_cert_of_pos pos = +let q_cert_of_pos pos = let rec _cert_of_pos = function Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i) | Axiom_le i -> Mc.PsatzIn (Ml2C.nat i) @@ -651,7 +651,7 @@ let rec term_to_z_expr = function let term_to_z_pol e = Mc.norm_aux (Ml2C.z 0) (Ml2C.z 1) Mc.Z.add Mc.Z.mul Mc.Z.sub Mc.Z.opp Mc.zeq_bool (term_to_z_expr e) -let z_cert_of_pos pos = +let z_cert_of_pos pos = let s,pos = (scale_certificate pos) in let rec _cert_of_pos = function Axiom_eq i -> Mc.PsatzIn (Ml2C.nat i) @@ -689,7 +689,7 @@ type prf_sys = (cstr * ProofFormat.prf_rule) list (** Proof generating pivoting over variable v *) -let pivot v (c1,p1) (c2,p2) = +let pivot v (c1,p1) (c2,p2) = let {coeffs = v1 ; op = op1 ; cst = n1} = c1 and {coeffs = v2 ; op = op2 ; cst = n2} = c2 in @@ -726,7 +726,7 @@ let pivot v (c1,p1) (c2,p2) = else None (* op2 could be Eq ... this might happen *) -let simpl_sys sys = +let simpl_sys sys = List.fold_left (fun acc (c,p) -> match check_int_sat (c,p) with | Tauto -> acc @@ -739,7 +739,7 @@ let simpl_sys sys = [ext_gcd a b = (x,y,g)] iff [ax+by=g] Source: http://en.wikipedia.org/wiki/Extended_Euclidean_algorithm *) -let rec ext_gcd a b = +let rec ext_gcd a b = if Int.equal (sign_big_int b) 0 then (unit_big_int,zero_big_int) else @@ -747,7 +747,7 @@ let rec ext_gcd a b = let (s,t) = ext_gcd b r in (t, sub_big_int s (mult_big_int q t)) -let extract_coprime (c1,p1) (c2,p2) = +let extract_coprime (c1,p1) (c2,p2) = if c1.op == Eq && c2.op == Eq then Vect.exists2 (fun n1 n2 -> Int.equal (compare_big_int (gcd_big_int (numerator n1) (numerator n2)) unit_big_int) 0) @@ -776,7 +776,7 @@ let extract_coprime_equation psys = let pivot_sys v pc psys = apply_and_normalise check_int_sat (pivot v pc) psys -let reduce_coprime psys = +let reduce_coprime psys = let oeq,sys = extract_coprime_equation psys in match oeq with | None -> None (* Nothing to do *) @@ -793,7 +793,7 @@ let reduce_coprime psys = Some (pivot_sys v (cstr,prf) ((c1,p1)::sys)) (** If there is an equation [eq] of the form 1.x + e = c, do a pivot over x with equation [eq] *) -let reduce_unary psys = +let reduce_unary psys = let is_unary_equation (cstr,prf) = if cstr.op == Eq then @@ -807,7 +807,7 @@ let reduce_unary psys = Some(pivot_sys v pc sys) -let reduce_var_change psys = +let reduce_var_change psys = let rec rel_prime vect = match Vect.choose vect with @@ -854,7 +854,7 @@ let reduction_equations psys = (** [get_bound sys] returns upon success an interval (lb,e,ub) with proofs *) -let get_bound sys = +let get_bound sys = let is_small (v,i) = match Itv.range i with | None -> false @@ -909,12 +909,12 @@ let get_bound sys = | None -> None -let check_sys sys = +let check_sys sys = List.for_all (fun (c,p) -> Vect.for_all (fun _ n -> sign_num n <> 0) c.coeffs) sys open ProofFormat -let xlia (can_enum:bool) reduction_equations sys = +let xlia (can_enum:bool) reduction_equations sys = let rec enum_proof (id:int) (sys:prf_sys) = @@ -979,9 +979,9 @@ let xlia (can_enum:bool) reduction_equations sys = end; let prf = compile_proof env prf in (*try - if Mc.zChecker sys' prf then Some prf else - raise Certificate.BadCertificate - with Failure s -> (Printf.printf "%s" s ; Some prf) + if Mc.zChecker sys' prf then Some prf else + raise Certificate.BadCertificate + with Failure s -> (Printf.printf "%s" s ; Some prf) *) Prf prf let xlia_simplex env red sys = @@ -1029,7 +1029,7 @@ let gen_bench (tac, prover) can_enum prfdepth sys = end); res -let lia (can_enum:bool) (prfdepth:int) sys = +let lia (can_enum:bool) (prfdepth:int) sys = let sys = develop_constraints prfdepth z_spec sys in if debug then begin Printf.fprintf stdout "Input problem\n"; @@ -1049,7 +1049,7 @@ let lia (can_enum:bool) (prfdepth:int) sys = let make_cstr_system sys = List.map (fun ((p,o),prf) -> (cstr_of_poly (p,o), prf)) sys -let nlia enum prfdepth sys = +let nlia enum prfdepth sys = let sys = develop_constraints prfdepth z_spec sys in let is_linear = List.for_all (fun ((p,_),_) -> LinPoly.is_linear p) sys in diff --git a/plugins/micromega/csdpcert.ml b/plugins/micromega/csdpcert.ml index cf5f60fb55..09e354957a 100644 --- a/plugins/micromega/csdpcert.ml +++ b/plugins/micromega/csdpcert.ml @@ -62,9 +62,9 @@ let partition_expr l = | Mc.Equal -> ((e,i)::eq,ge,neq) | Mc.NonStrict -> (eq,(e,Axiom_le i)::ge,neq) | Mc.Strict -> (* e > 0 == e >= 0 /\ e <> 0 *) - (eq, (e,Axiom_lt i)::ge,(e,Axiom_lt i)::neq) + (eq, (e,Axiom_lt i)::ge,(e,Axiom_lt i)::neq) | Mc.NonEqual -> (eq,ge,(e,Axiom_eq i)::neq) - (* Not quite sure -- Coq interface has changed *) + (* Not quite sure -- Coq interface has changed *) in f 0 l @@ -72,7 +72,7 @@ let rec sets_of_list l = match l with | [] -> [[]] | e::l -> let s = sets_of_list l in - s@(List.map (fun s0 -> e::s0) s) + s@(List.map (fun s0 -> e::s0) s) (* The exploration is probably not complete - for simple cases, it works... *) let real_nonlinear_prover d l = @@ -83,9 +83,9 @@ let real_nonlinear_prover d l = let rec elim_const = function [] -> [] | (x,y)::l -> let p = poly_of_term (expr_to_term x) in - if poly_isconst p - then elim_const l - else (p,y)::(elim_const l) in + if poly_isconst p + then elim_const l + else (p,y)::(elim_const l) in let eq = elim_const eq in let peq = List.map fst eq in @@ -104,7 +104,7 @@ let real_nonlinear_prover d l = let (cert_ideal, cert_cone,monoid) = deepen_until d (fun d -> tryfind (fun m -> let (ci,cc) = real_positivnullstellensatz_general false d peq pge (poly_neg (fst m) ) in - (ci,cc,snd m)) monoids) 0 in + (ci,cc,snd m)) monoids) 0 in let proofs_ideal = List.map2 (fun q i -> Eqmul(term_of_poly q,Axiom_eq i)) cert_ideal (List.map snd eq) in @@ -141,9 +141,9 @@ let pure_sos l = let plt = poly_neg (poly_of_term (expr_to_term (fst lt))) in let (n,polys) = sumofsquares plt in (* n * (ci * pi^2) *) let pos = Product (Rational_lt n, - List.fold_right (fun (c,p) rst -> Sum (Product (Rational_lt c, Square - (term_of_poly p)), rst)) - polys (Rational_lt (Int 0))) in + List.fold_right (fun (c,p) rst -> Sum (Product (Rational_lt c, Square + (term_of_poly p)), rst)) + polys (Rational_lt (Int 0))) in let proof = Sum(Axiom_lt i, pos) in (* let s,proof' = scale_certificate proof in let cert = snd (cert_of_pos proof') in *) diff --git a/plugins/micromega/g_micromega.mlg b/plugins/micromega/g_micromega.mlg index bcf546f059..edf8106f30 100644 --- a/plugins/micromega/g_micromega.mlg +++ b/plugins/micromega/g_micromega.mlg @@ -56,7 +56,7 @@ TACTIC EXTEND NQA END - + TACTIC EXTEND Sos_Z | [ "sos_Z" tactic(t) ] -> { (Coq_micromega.sos_Z (Tacinterp.tactic_of_value ist t)) } END diff --git a/plugins/micromega/mfourier.ml b/plugins/micromega/mfourier.ml index 943bcb384b..75cdfa24f1 100644 --- a/plugins/micromega/mfourier.ml +++ b/plugins/micromega/mfourier.ml @@ -71,13 +71,13 @@ exception SystemContradiction of proof let pp_cstr o (vect,bnd) = let (l,r) = bnd in (match l with - | None -> () - | Some n -> Printf.fprintf o "%s <= " (string_of_num n)) + | None -> () + | Some n -> Printf.fprintf o "%s <= " (string_of_num n)) ; Vect.pp o vect ; (match r with - | None -> output_string o"\n" - | Some n -> Printf.fprintf o "<=%s\n" (string_of_num n)) + | None -> output_string o"\n" + | Some n -> Printf.fprintf o "<=%s\n" (string_of_num n)) let pp_system o sys= @@ -96,7 +96,7 @@ let merge_cstr_info i1 i2 = match inter i1 i2 with | None -> None (* Could directly raise a system contradiction exception *) | Some bnd -> - Some { pos = p1 ; neg = n1 ; bound = bnd ; prf = And(prf1,prf2) } + Some { pos = p1 ; neg = n1 ; bound = bnd ; prf = And(prf1,prf2) } (** [xadd_cstr vect cstr_info] loads an constraint into the system. The constraint is neither redundant nor contradictory. @@ -107,14 +107,14 @@ let xadd_cstr vect cstr_info sys = try let info = System.find sys vect in match merge_cstr_info cstr_info !info with - | None -> raise (SystemContradiction (And(cstr_info.prf, (!info).prf))) - | Some info' -> info := info' + | None -> raise (SystemContradiction (And(cstr_info.prf, (!info).prf))) + | Some info' -> info := info' with - | Not_found -> System.replace sys vect (ref cstr_info) + | Not_found -> System.replace sys vect (ref cstr_info) exception TimeOut - -let xadd_cstr vect cstr_info sys = + +let xadd_cstr vect cstr_info sys = if debug && Int.equal (System.length sys mod 1000) 0 then (print_string "*" ; flush stdout) ; if System.length sys < !max_nb_cstr then xadd_cstr vect cstr_info sys @@ -122,11 +122,11 @@ let xadd_cstr vect cstr_info sys = type cstr_ext = | Contradiction (** The constraint is contradictory. - Typically, a [SystemContradiction] exception will be raised. *) + Typically, a [SystemContradiction] exception will be raised. *) | Redundant (** The constrain is redundant. - Typically, the constraint will be dropped *) + Typically, the constraint will be dropped *) | Cstr of vector * cstr_info (** Taken alone, the constraint is neither contradictory nor redundant. - Typically, it will be added to the constraint system. *) + Typically, it will be added to the constraint system. *) (** [normalise_cstr] : vector -> cstr_info -> cstr_ext *) let normalise_cstr vect cinfo = @@ -136,8 +136,8 @@ let normalise_cstr vect cinfo = match Vect.choose vect with | None -> if Itv.in_bound (l,r) (Int 0) then Redundant else Contradiction | Some (_,n,_) -> Cstr(Vect.div n vect, - let divn x = x // n in - if Int.equal (sign_num n) 1 + let divn x = x // n in + if Int.equal (sign_num n) 1 then{cinfo with bound = (Option.map divn l , Option.map divn r) } else {cinfo with pos = cinfo.neg ; neg = cinfo.pos ; bound = (Option.map divn r , Option.map divn l)}) @@ -157,11 +157,11 @@ let norm_cstr {coeffs = v ; op = o ; cst = c} idx = normalise_cstr v {pos = p ; neg = n ; bound = (match o with - | Eq -> Some c , Some c + | Eq -> Some c , Some c | Ge -> Some c , None | Gt -> raise Polynomial.Strict ) ; - prf = Assum idx } + prf = Assum idx } (** [load_system l] takes a list of constraints of type [cstr_compat] @@ -179,7 +179,7 @@ let load_system l = | Contradiction -> raise (SystemContradiction (Assum i)) | Redundant -> vrs | Cstr(vect,info) -> - xadd_cstr vect info sys ; + xadd_cstr vect info sys ; Vect.fold (fun s v _ -> ISet.add v s) vrs cstr.coeffs) ISet.empty li in {sys = sys ;vars = vars} @@ -218,7 +218,7 @@ let add (v1,c1) (v2,c2) = let split x (vect: vector) info (l,m,r) = match get x vect with | Int 0 -> (* The constraint does not mention [x], store it in m *) - (l,(vect,info)::m,r) + (l,(vect,info)::m,r) | vl -> (* otherwise *) let cons_bound lst bd = @@ -257,10 +257,10 @@ let project vr sys = List.iter(fun l_elem -> List.iter (fun r_elem -> let (vect,info) = elim l_elem r_elem in - match normalise_cstr vect info with - | Redundant -> () - | Contradiction -> raise (SystemContradiction info.prf) - | Cstr(vect,info) -> xadd_cstr vect info new_sys) r ) l; + match normalise_cstr vect info with + | Redundant -> () + | Contradiction -> raise (SystemContradiction info.prf) + | Cstr(vect,info) -> xadd_cstr vect info new_sys) r ) l; {sys = new_sys ; vars = ISet.remove vr sys.vars} @@ -277,20 +277,20 @@ let project_using_eq vr c vect bound prf (vect',info') = match get vr vect' with | Int 0 -> (vect',info') | c2 -> - let c1 = if c2 >=/ Int 0 then minus_num c else c in + let c1 = if c2 >=/ Int 0 then minus_num c else c in - let c2 = abs_num c2 in + let c2 = abs_num c2 in - let (vres,(n,p)) = add (vect,c1) (vect', c2) in + let (vres,(n,p)) = add (vect,c1) (vect', c2) in - let cst = bound // c1 in + let cst = bound // c1 in - let bndres = - let f x = cst +/ x // c2 in - let (l,r) = info'.bound in + let bndres = + let f x = cst +/ x // c2 in + let (l,r) = info'.bound in (Option.map f l , Option.map f r) in - (vres,{neg = n ; pos = p ; bound = bndres ; prf = Elim(vr,prf,info'.prf)}) + (vres,{neg = n ; pos = p ; bound = bndres ; prf = Elim(vr,prf,info'.prf)}) let elim_var_using_eq vr vect cst prf sys = @@ -302,10 +302,10 @@ let elim_var_using_eq vr vect cst prf sys = System.iter(fun vect iref -> let (vect',info') = elim_var (vect,!iref) in - match normalise_cstr vect' info' with - | Redundant -> () - | Contradiction -> raise (SystemContradiction info'.prf) - | Cstr(vect,info') -> xadd_cstr vect info' new_sys) sys.sys ; + match normalise_cstr vect' info' with + | Redundant -> () + | Contradiction -> raise (SystemContradiction info'.prf) + | Cstr(vect,info') -> xadd_cstr vect info' new_sys) sys.sys ; {sys = new_sys ; vars = ISet.remove vr sys.vars} @@ -337,8 +337,8 @@ let restrict_bound n sum (itv:interval) = let l,r = itv in match sign_num n with | 0 -> if in_bound itv sum - then (None,None) (* redundant *) - else failwith "SystemContradiction" + then (None,None) (* redundant *) + else failwith "SystemContradiction" | 1 -> Option.map f l , Option.map f r | _ -> Option.map f r , Option.map f l @@ -355,7 +355,7 @@ let bound_of_variable map v sys = Vect.pp vect (Num.string_of_num sum) Vect.pp rst ; Printf.fprintf stdout "current interval: %a\n" Itv.pp (!iref).bound; failwith "bound_of_variable: impossible" - | Some itv -> itv) sys (None,None) + | Some itv -> itv) sys (None,None) (** [pick_small_value bnd] picks a value being closed to zero within the interval *) @@ -365,10 +365,10 @@ let pick_small_value bnd = | None , Some i -> if (Int 0) <=/ (floor_num i) then Int 0 else floor_num i | Some i,None -> if i <=/ (Int 0) then Int 0 else ceiling_num i | Some i,Some j -> - if i <=/ Int 0 && Int 0 <=/ j - then Int 0 - else if ceiling_num i <=/ floor_num j - then ceiling_num i (* why not *) else i + if i <=/ Int 0 && Int 0 <=/ j + then Int 0 + else if ceiling_num i <=/ floor_num j + then ceiling_num i (* why not *) else i (** [solution s1 sys_l = Some(sn,\[(vn-1,sn-1);...; (v1,s1)\]\@sys_l)] @@ -385,20 +385,20 @@ let solve_sys black_v choose_eq choose_variable sys sys_l = let eqs = choose_eq sys in try - let (v,vect,cst,ln) = fst (List.find (fun ((v,_,_,_),_) -> v <> black_v) eqs) in - if debug then + let (v,vect,cst,ln) = fst (List.find (fun ((v,_,_,_),_) -> v <> black_v) eqs) in + if debug then (Printf.printf "\nE %a = %s variable %i\n" Vect.pp vect (string_of_num cst) v ; - flush stdout); - let sys' = elim_var_using_eq v vect cst ln sys in - solve_sys sys' ((v,sys)::sys_l) + flush stdout); + let sys' = elim_var_using_eq v vect cst ln sys in + solve_sys sys' ((v,sys)::sys_l) with Not_found -> let vars = choose_variable sys in - try - let (v,est) = (List.find (fun (v,_) -> v <> black_v) vars) in - if debug then (Printf.printf "\nV : %i estimate %f\n" v est ; flush stdout) ; - let sys' = project v sys in - solve_sys sys' ((v,sys)::sys_l) - with Not_found -> (* we are done *) Inl (sys,sys_l) in + try + let (v,est) = (List.find (fun (v,_) -> v <> black_v) vars) in + if debug then (Printf.printf "\nV : %i estimate %f\n" v est ; flush stdout) ; + let sys' = project v sys in + solve_sys sys' ((v,sys)::sys_l) + with Not_found -> (* we are done *) Inl (sys,sys_l) in solve_sys sys sys_l @@ -408,7 +408,7 @@ let solve black_v choose_eq choose_variable cstrs = try let sys = load_system cstrs in - if debug then Printf.printf "solve :\n %a" pp_system sys.sys ; + if debug then Printf.printf "solve :\n %a" pp_system sys.sys ; solve_sys black_v choose_eq choose_variable sys [] with SystemContradiction prf -> Inr prf @@ -430,20 +430,20 @@ struct match Vect.choose l1 with | None -> xpart rl ((Vect.null,info)::ltl) n (info.neg+info.pos+z) p | Some(vr, vl, rl1) -> - if Int.equal v vr - then - let cons_bound lst bd = - match bd with - | None -> lst - | Some bnd -> info.neg+info.pos::lst in - - let lb,rb = info.bound in - if Int.equal (sign_num vl) 1 - then xpart rl ((rl1,info)::ltl) (cons_bound n lb) z (cons_bound p rb) - else xpart rl ((rl1,info)::ltl) (cons_bound n rb) z (cons_bound p lb) - else - (* the variable is greater *) - xpart rl ((l1,info)::ltl) n (info.neg+info.pos+z) p + if Int.equal v vr + then + let cons_bound lst bd = + match bd with + | None -> lst + | Some bnd -> info.neg+info.pos::lst in + + let lb,rb = info.bound in + if Int.equal (sign_num vl) 1 + then xpart rl ((rl1,info)::ltl) (cons_bound n lb) z (cons_bound p rb) + else xpart rl ((rl1,info)::ltl) (cons_bound n rb) z (cons_bound p lb) + else + (* the variable is greater *) + xpart rl ((l1,info)::ltl) n (info.neg+info.pos+z) p in let (sys',n,z,p) = xpart l [] [] 0 [] in @@ -484,15 +484,15 @@ struct match Vect.choose l with | None -> (false,Vect.null) | Some(i,_,rl) -> if Int.equal i v - then (true,rl) - else if i < v then unroll_until v rl else (false,l) + then (true,rl) + else if i < v then unroll_until v rl else (false,l) - let rec choose_simple_equation eqs = + let rec choose_simple_equation eqs = match eqs with | [] -> None - | (vect,a,prf,ln)::eqs -> + | (vect,a,prf,ln)::eqs -> match Vect.choose vect with | Some(i,v,rst) -> if Vect.is_null rst then Some (i,vect,a,prf,ln) @@ -507,29 +507,29 @@ struct *) let is_primal_equation_var v = List.fold_left (fun nb_eq (vect,info) -> - if fst (unroll_until v vect) - then if itv_point info.bound then nb_eq + 1 else nb_eq - else nb_eq) 0 sys_l in + if fst (unroll_until v vect) + then if itv_point info.bound then nb_eq + 1 else nb_eq + else nb_eq) 0 sys_l in let rec find_var vect = match Vect.choose vect with | None -> None | Some(i,_,vect) -> - let nb_eq = is_primal_equation_var i in - if Int.equal nb_eq 2 - then Some i else find_var vect in + let nb_eq = is_primal_equation_var i in + if Int.equal nb_eq 2 + then Some i else find_var vect in let rec find_eq_var eqs = match eqs with - | [] -> None - | (vect,a,prf,ln)::l -> - match find_var vect with - | None -> find_eq_var l - | Some r -> Some (r,vect,a,prf,ln) + | [] -> None + | (vect,a,prf,ln)::l -> + match find_var vect with + | None -> find_eq_var l + | Some r -> Some (r,vect,a,prf,ln) in match choose_simple_equation eqs with - | None -> find_eq_var eqs - | Some res -> Some res + | None -> find_eq_var eqs + | Some res -> Some res @@ -539,43 +539,43 @@ struct let equalities = List.fold_left (fun l (vect,info) -> - match info.bound with - | Some a , Some b -> - if a =/ b then (* This an equation *) - (vect,a,info.prf,info.neg+info.pos)::l else l - | _ -> l + match info.bound with + | Some a , Some b -> + if a =/ b then (* This an equation *) + (vect,a,info.prf,info.neg+info.pos)::l else l + | _ -> l ) [] sys_l in let rec estimate_cost v ct sysl acc tlsys = match sysl with - | [] -> (acc,tlsys) - | (l,info)::rsys -> - let ln = info.pos + info.neg in - let (b,l) = unroll_until v l in - match b with - | true -> - if itv_point info.bound - then estimate_cost v ct rsys (acc+ln) ((l,info)::tlsys) (* this is free *) - else estimate_cost v ct rsys (acc+ln+ct) ((l,info)::tlsys) (* should be more ? *) - | false -> estimate_cost v ct rsys (acc+ln) ((l,info)::tlsys) in + | [] -> (acc,tlsys) + | (l,info)::rsys -> + let ln = info.pos + info.neg in + let (b,l) = unroll_until v l in + match b with + | true -> + if itv_point info.bound + then estimate_cost v ct rsys (acc+ln) ((l,info)::tlsys) (* this is free *) + else estimate_cost v ct rsys (acc+ln+ct) ((l,info)::tlsys) (* should be more ? *) + | false -> estimate_cost v ct rsys (acc+ln) ((l,info)::tlsys) in match choose_primal_equation equalities sys_l with - | None -> - let cost_eq eq const prf ln acc_costs = + | None -> + let cost_eq eq const prf ln acc_costs = - let rec cost_eq eqr sysl costs = + let rec cost_eq eqr sysl costs = match Vect.choose eqr with | None -> costs | Some(v,_,eqr) -> let (cst,tlsys) = estimate_cost v (ln-1) sysl 0 [] in - cost_eq eqr tlsys (((v,eq,const,prf),cst)::costs) in - cost_eq eq sys_l acc_costs in + cost_eq eqr tlsys (((v,eq,const,prf),cst)::costs) in + cost_eq eq sys_l acc_costs in - let all_costs = List.fold_left (fun all_costs (vect,const,prf,ln) -> cost_eq vect const prf ln all_costs) [] equalities in + let all_costs = List.fold_left (fun all_costs (vect,const,prf,ln) -> cost_eq vect const prf ln all_costs) [] equalities in - (* pp_list (fun o ((v,eq,_,_),cst) -> Printf.fprintf o "((%i,%a),%i)\n" v pp_vect eq cst) stdout all_costs ; *) + (* pp_list (fun o ((v,eq,_,_),cst) -> Printf.fprintf o "((%i,%a),%i)\n" v pp_vect eq cst) stdout all_costs ; *) - List.sort (fun x y -> Int.compare (snd x) (snd y) ) all_costs - | Some (v,vect, const,prf,_) -> [(v,vect,const,prf),0] + List.sort (fun x y -> Int.compare (snd x) (snd y) ) all_costs + | Some (v,vect, const,prf,_) -> [(v,vect,const,prf),0] end @@ -593,12 +593,12 @@ struct op = Eq ; cst = (Int 0)} in match solve fresh choose_equality_var choose_variable (cstr::l) with - | Inr prf -> None (* This is an unsatisfiability proof *) - | Inl (s,_) -> - try - Some (bound_of_variable IMap.empty fresh s.sys) - with x when CErrors.noncritical x -> - Printf.printf "optimise Exception : %s" (Printexc.to_string x); + | Inr prf -> None (* This is an unsatisfiability proof *) + | Inl (s,_) -> + try + Some (bound_of_variable IMap.empty fresh s.sys) + with x when CErrors.noncritical x -> + Printf.printf "optimise Exception : %s" (Printexc.to_string x); None @@ -608,16 +608,16 @@ struct | Inr prf -> Inr prf | Inl (_,l) -> - let rec rebuild_solution l map = - match l with - | [] -> map - | (v,e)::l -> - let itv = bound_of_variable map v e.sys in - let map = IMap.add v (pick_small_value itv) map in - rebuild_solution l map - in + let rec rebuild_solution l map = + match l with + | [] -> map + | (v,e)::l -> + let itv = bound_of_variable map v e.sys in + let map = IMap.add v (pick_small_value itv) map in + rebuild_solution l map + in - let map = rebuild_solution l IMap.empty in + let map = rebuild_solution l IMap.empty in let vect = IMap.fold (fun v i vect -> Vect.set v i vect) map Vect.null in if debug then Printf.printf "SOLUTION %a" Vect.pp vect ; let res = Inl vect in @@ -645,9 +645,9 @@ struct let forall_pairs f l1 l2 = List.fold_left (fun acc e1 -> List.fold_left (fun acc e2 -> - match f e1 e2 with - | None -> acc - | Some v -> v::acc) acc l2) [] l1 + match f e1 e2 with + | None -> acc + | Some v -> v::acc) acc l2) [] l1 let add_op x y = @@ -664,8 +664,8 @@ struct | Int 0 , _ | _ , Int 0 -> None | a , b -> if Int.equal ((sign_num a) * (sign_num b)) (-1) - then - Some (add (p1,abs_num a) (p2,abs_num b) , + then + Some (add (p1,abs_num a) (p2,abs_num b) , {coeffs = add (v1,abs_num a) (v2,abs_num b) ; op = add_op op1 op2 ; cst = n1 // (abs_num a) +/ n2 // (abs_num b) }) @@ -675,12 +675,12 @@ struct op = add_op op1 op2; cst = n1 // (minus_num (a// b)) +/ n2 // (Int 1)}) else if op2 == Eq - then - Some (add (p2,minus_num (b // a)) (p1,Int 1), + then + Some (add (p2,minus_num (b // a)) (p1,Int 1), {coeffs = add (v2,minus_num (b// a)) (v1 ,Int 1) ; op = add_op op1 op2; cst = n2 // (minus_num (b// a)) +/ n1 // (Int 1)}) - else None (* op2 could be Eq ... this might happen *) + else None (* op2 could be Eq ... this might happen *) let normalise_proofs l = @@ -752,10 +752,10 @@ let mk_proof hyps prf = match prfs with | Inr x -> [x] | Inl (oleft,oright) -> - match oleft , oright with - | None , None -> [] - | None , Some(prf,cstr,_) | Some(prf,cstr,_) , None -> [prf,cstr] - | Some(prf1,cstr1,_) , Some(prf2,cstr2,_) -> [prf1,cstr1;prf2,cstr2] in + match oleft , oright with + | None , None -> [] + | None , Some(prf,cstr,_) | Some(prf,cstr,_) , None -> [prf,cstr] + | Some(prf1,cstr1,_) , Some(prf2,cstr2,_) -> [prf1,cstr1;prf2,cstr2] in mk_proof prf diff --git a/plugins/micromega/mutils.ml b/plugins/micromega/mutils.ml index cca66c0719..a30e963f2a 100644 --- a/plugins/micromega/mutils.ml +++ b/plugins/micromega/mutils.ml @@ -66,7 +66,7 @@ let rec try_any l x = let all_pairs f l = let pair_with acc e l = List.fold_left (fun acc x -> (f e x) ::acc) acc l in - let rec xpairs acc l = + let rec xpairs acc l = match l with | [] -> acc | e::lx -> xpairs (pair_with acc e l) lx in @@ -77,20 +77,20 @@ let rec is_sublist f l1 l2 = | [] ,_ -> true | e::l1', [] -> false | e::l1' , e'::l2' -> - if f e e' then is_sublist f l1' l2' - else is_sublist f l1 l2' - -let extract pred l = - List.fold_left (fun (fd,sys) e -> - match fd with - | None -> - begin - match pred e with - | None -> fd, e::sys - | Some v -> Some(v,e) , sys - end - | _ -> (fd, e::sys) - ) (None,[]) l + if f e e' then is_sublist f l1' l2' + else is_sublist f l1 l2' + +let extract pred l = + List.fold_left (fun (fd,sys) e -> + match fd with + | None -> + begin + match pred e with + | None -> fd, e::sys + | Some v -> Some(v,e) , sys + end + | _ -> (fd, e::sys) + ) (None,[]) l let extract_best red lt l = let rec extractb c e rst l = @@ -338,7 +338,7 @@ struct end (** - * MODULE: Labels for atoms in propositional formulas. + * MODULE: Labels for atoms in propositional formulas. * Tags are used to identify unused atoms in CNFs, and propagate them back to * the original formula. The translation back to Coq then ignores these * superfluous items, which speeds the translation up a bit. @@ -406,26 +406,26 @@ let command exe_path args vl = finally (* Recover the result *) - (fun () -> - match status with - | Unix.WEXITED 0 -> - let inch = Unix.in_channel_of_descr stdout_read in - begin + (fun () -> + match status with + | Unix.WEXITED 0 -> + let inch = Unix.in_channel_of_descr stdout_read in + begin try Marshal.from_channel inch with any -> failwith (Printf.sprintf "command \"%s\" exited %s" exe_path (Printexc.to_string any)) end - | Unix.WEXITED i -> + | Unix.WEXITED i -> failwith (Printf.sprintf "command \"%s\" exited %i" exe_path i) - | Unix.WSIGNALED i -> + | Unix.WSIGNALED i -> failwith (Printf.sprintf "command \"%s\" killed %i" exe_path i) - | Unix.WSTOPPED i -> + | Unix.WSTOPPED i -> failwith (Printf.sprintf "command \"%s\" stopped %i" exe_path i)) (* Cleanup *) - (fun () -> - List.iter (fun x -> try Unix.close x with any -> ()) + (fun () -> + List.iter (fun x -> try Unix.close x with any -> ()) [stdin_read; stdin_write; stdout_read; stdout_write; stderr_read; stderr_write]) diff --git a/plugins/micromega/persistent_cache.ml b/plugins/micromega/persistent_cache.ml index 14e2e40846..28d8d5a020 100644 --- a/plugins/micromega/persistent_cache.ml +++ b/plugins/micromega/persistent_cache.ml @@ -49,9 +49,9 @@ struct type 'a t = { - outch : out_channel ; - mutable status : mode ; - htbl : 'a Table.t + outch : out_channel ; + mutable status : mode ; + htbl : 'a Table.t } @@ -72,49 +72,49 @@ let read_key_elem inch = | End_of_file -> None | e when CErrors.noncritical e -> raise InvalidTableFormat -(** +(** We used to only lock/unlock regions. - Is-it more robust/portable to lock/unlock a fixed region e.g. [0;1]? + Is-it more robust/portable to lock/unlock a fixed region e.g. [0;1]? In case of locking failure, the cache is not used. **) type lock_kind = Read | Write -let lock kd fd = - let pos = lseek fd 0 SEEK_CUR in - let success = - try +let lock kd fd = + let pos = lseek fd 0 SEEK_CUR in + let success = + try ignore (lseek fd 0 SEEK_SET); - let lk = match kd with - | Read -> F_RLOCK - | Write -> F_LOCK in + let lk = match kd with + | Read -> F_RLOCK + | Write -> F_LOCK in lockf fd lk 1; true with Unix.Unix_error(_,_,_) -> false in - ignore (lseek fd pos SEEK_SET) ; + ignore (lseek fd pos SEEK_SET) ; success -let unlock fd = +let unlock fd = let pos = lseek fd 0 SEEK_CUR in - try - ignore (lseek fd 0 SEEK_SET) ; + try + ignore (lseek fd 0 SEEK_SET) ; lockf fd F_ULOCK 1 - with - Unix.Unix_error(_,_,_) -> () - (* Here, this is really bad news -- + with + Unix.Unix_error(_,_,_) -> () + (* Here, this is really bad news -- there is a pending lock which could cause a deadlock. Should it be an anomaly or produce a warning ? *); - ignore (lseek fd pos SEEK_SET) + ignore (lseek fd pos SEEK_SET) (* We make the assumption that an acquired lock can always be released *) -let do_under_lock kd fd f = +let do_under_lock kd fd f = if lock kd fd then finally f (fun () -> unlock fd) else f () - + let open_in f = @@ -128,11 +128,11 @@ let open_in f = | None -> () | Some (key,elem) -> Table.add htbl key elem ; - xload () in + xload () in try (* Locking of the (whole) file while reading *) - do_under_lock Read finch xload ; - close_in_noerr inch ; + do_under_lock Read finch xload ; + close_in_noerr inch ; { outch = out_channel_of_descr (openfile f [O_WRONLY;O_APPEND;O_CREAT] 0o666) ; status = Open ; @@ -145,11 +145,11 @@ let open_in f = let flags = [O_WRONLY; O_TRUNC;O_CREAT] in let out = (openfile f flags 0o666) in let outch = out_channel_of_descr out in - do_under_lock Write out - (fun () -> - Table.iter - (fun k e -> Marshal.to_channel outch (k,e) [Marshal.No_sharing]) htbl; - flush outch) ; + do_under_lock Write out + (fun () -> + Table.iter + (fun k e -> Marshal.to_channel outch (k,e) [Marshal.No_sharing]) htbl; + flush outch) ; { outch = outch ; status = Open ; htbl = htbl @@ -165,8 +165,8 @@ let add t k e = let fd = descr_of_out_channel outch in begin Table.add tbl k e ; - do_under_lock Write fd - (fun _ -> + do_under_lock Write fd + (fun _ -> Marshal.to_channel outch (k,e) [Marshal.No_sharing] ; flush outch ) @@ -178,7 +178,7 @@ let find t k = then raise UnboundTable else let res = Table.find tbl k in - res + res let memo cache f = let tbl = lazy (try Some (open_in cache) with _ -> None) in @@ -186,13 +186,13 @@ let memo cache f = match Lazy.force tbl with | None -> f x | Some tbl -> - try - find tbl x - with - Not_found -> - let res = f x in - add tbl x res ; - res + try + find tbl x + with + Not_found -> + let res = f x in + add tbl x res ; + res let memo_cond cache cond f = let tbl = lazy (try Some (open_in cache) with _ -> None) in diff --git a/plugins/micromega/sos_lib.ml b/plugins/micromega/sos_lib.ml index 58d5d7ecf1..0a0ffc7947 100644 --- a/plugins/micromega/sos_lib.ml +++ b/plugins/micromega/sos_lib.ml @@ -525,11 +525,11 @@ let deepen_until limit f n = | 0 -> raise TooDeep | -1 -> deepen f n | _ -> - let rec d_until f n = - try(* if !debugging - then (print_string "Searching with depth limit "; - print_int n; print_newline()) ;*) f n - with Failure x -> - (*if !debugging then (Printf.printf "solver error : %s\n" x) ; *) - if n = limit then raise TooDeep else d_until f (n + 1) in - d_until f n + let rec d_until f n = + try(* if !debugging + then (print_string "Searching with depth limit "; + print_int n; print_newline()) ;*) f n + with Failure x -> + (*if !debugging then (Printf.printf "solver error : %s\n" x) ; *) + if n = limit then raise TooDeep else d_until f (n + 1) in + d_until f n diff --git a/plugins/nsatz/ideal.ml b/plugins/nsatz/ideal.ml index 7ea56b41ec..46685e6a63 100644 --- a/plugins/nsatz/ideal.ml +++ b/plugins/nsatz/ideal.ml @@ -87,13 +87,13 @@ let compare_mon (m : mon) (m' : mon) = (* degre lexicographique inverse *) match Int.compare m.(0) m'.(0) with | 0 -> (* meme degre total *) - let res=ref 0 in - let i=ref d in - while (!res=0) && (!i>=1) do - res:= - (Int.compare m.(!i) m'.(!i)); - i:=!i-1; - done; - !res + let res=ref 0 in + let i=ref d in + while (!res=0) && (!i>=1) do + res:= - (Int.compare m.(!i) m'.(!i)); + i:=!i-1; + done; + !res | x -> x) let div_mon m m' = @@ -135,13 +135,13 @@ let ppcm_mon m m' = (* returns a constant polynom ial with d variables *) let const_mon d = let m = Array.make (d+1) 0 in - let m = set_deg m in + let m = set_deg m in m let var_mon d i = let m = Array.make (d+1) 0 in m.(i) <- 1; - let m = set_deg m in + let m = set_deg m in m end @@ -174,7 +174,7 @@ type polynom = { (********************************************************************** Polynomials - list of (coefficient, monomial) decreasing order + list of (coefficient, monomial) decreasing order *) let repr p = p @@ -216,10 +216,10 @@ let string_of_pol zeroP hdP tlP coefterm monterm string_of_coef | e -> s:= (!s) @ [((getvar lvar (i-1)) ^ "^" ^ e)]); done; (match !s with - [] -> if coefone + [] -> if coefone then "1" else "" - | l -> if coefone + | l -> if coefone then (String.concat "*" l) else ( "*" ^ (String.concat "*" l))) @@ -233,26 +233,26 @@ let string_of_pol zeroP hdP tlP coefterm monterm string_of_coef | "-1" ->( "-" ^" "^(string_of_mon m true)) | c -> if (String.get c 0)='-' then ( "- "^ - (String.sub c 1 + (String.sub c 1 ((String.length c)-1))^ (string_of_mon m false)) else (match start with true -> ( c^(string_of_mon m false)) |false -> ( "+ "^ c^(string_of_mon m false))) - and stringP p start = + and stringP p start = if (zeroP p) - then (if start + then (if start then ("0") else "") else ((string_of_term (hdP p) start)^ " "^ (stringP (tlP p) false)) - in + in (stringP p true) let stringP metadata (p : poly) = - string_of_pol + string_of_pol (fun p -> match p with [] -> true | _ -> false) (fun p -> match p with (t::p) -> t |_ -> failwith "print_pol dans dansideal") (fun p -> match p with (t::p) -> p |_ -> failwith "print_pol dans dansideal") @@ -309,7 +309,7 @@ let coef_of_int x = P.of_num (Num.Int x) (* variable i *) let gen d i = - let m = var_mon d i in + let m = var_mon d i in [((coef_of_int 1),m)] let oppP p = @@ -349,7 +349,7 @@ let puisP p n= let q = multP q q in if n mod 2 = 0 then q else multP p q in puisP p n - + (*********************************************************************** Division of polynomials *) @@ -366,7 +366,7 @@ type table = { let pgcdpos a b = P.pgcdP a b let polynom0 = { pol = []; num = 0 } - + let ppol p = p.pol let lm p = snd (List.hd (ppol p)) @@ -390,7 +390,7 @@ let rec selectdiv m l = true -> q |false -> selectdiv m r -let div_pol p q a b m = +let div_pol p q a b m = plusP (emultP a p) (mult_t_pol b m q) let find_hmon table m = match table.hmon with @@ -424,15 +424,15 @@ let reduce2 table p l = match p with [] -> (coef1,[]) |t::p' -> - let (a,m)=t in + let (a,m)=t in let q = selectdiv table m l in match q with - [] -> if reduire_les_queues - then - let (c,r)=(reduce p') in + [] -> if reduire_les_queues + then + let (c,r)=(reduce p') in (c,((P.multP a c,m)::r)) - else (coef1,p) - |(b,m')::q' -> + else (coef1,p) + |(b,m')::q' -> let c=(pgcdpos a b) in let a'= (div_coef b c) in let b'=(P.oppP (div_coef a c)) in @@ -450,7 +450,7 @@ let coefpoldep_find table p q = let coefpoldep_set table p q c = Hashtbl.add table.coefpoldep (p.num,q.num) c -(* keeps trace in coefpoldep +(* keeps trace in coefpoldep divides without pseudodivisions *) let reduce2_trace table p l lcp = @@ -461,16 +461,16 @@ let reduce2_trace table p l lcp = match p with [] -> ([],[]) |t::p' -> - let (a,m)=t in + let (a,m)=t in let q = selectdiv table m l in match q with - [] -> - if reduire_les_queues - then - let (lq,r)=(reduce p') in + [] -> + if reduire_les_queues + then + let (lq,r)=(reduce p') in (lq,((a,m)::r)) - else ([],p) - |(b,m')::q' -> + else ([],p) + |(b,m')::q' -> let b'=(P.oppP (div_coef a b)) in let m''= div_mon m m' in let p1=plusP p' (mult_t_pol b' m'' q') in @@ -480,18 +480,18 @@ let reduce2_trace table p l lcp = (List.map2 (fun c0 q -> let c = - List.fold_left - (fun x (a,m,s) -> - if equal s (ppol q) - then - plusP x (mult_t_pol a m (polconst (nvar m) (coef_of_int 1))) - else x) - c0 - lq in + List.fold_left + (fun x (a,m,s) -> + if equal s (ppol q) + then + plusP x (mult_t_pol a m (polconst (nvar m) (coef_of_int 1))) + else x) + c0 + lq in c) lcp lp, - r) + r) (*********************************************************************** Completion @@ -511,12 +511,12 @@ let spol0 ps qs= let m1 = div_mon m'' m in let m2 = div_mon m'' m' in let fsp p' q' = - plusP - (mult_t_pol - (div_coef b c) - m1 p') - (mult_t_pol - (P.oppP (div_coef a c)) + plusP + (mult_t_pol + (div_coef b c) + m1 p') + (mult_t_pol + (P.oppP (div_coef a c)) m2 q') in let sp = fsp p' q' in let p0 = fsp (polconst (nvar m) (coef_of_int 1)) [] in @@ -564,7 +564,7 @@ end let ord i j = if i<j then (i,j) else (j,i) - + let cpair p q accu = if etrangers (ppol p) (ppol q) then accu else Heap.add (ord p.num q.num, ppcm_mon (lm p) (lm q)) accu @@ -582,14 +582,14 @@ let critere3 table ((i,j),m) lp lcp = (fun h -> h.num <> i && h.num <> j && (div_mon_test m (lm h)) - && (h.num < j - || not (m = ppcm_mon - (lm (table.allpol.(i))) - (lm h))) - && (h.num < i - || not (m = ppcm_mon - (lm (table.allpol.(j))) - (lm h)))) + && (h.num < j + || not (m = ppcm_mon + (lm (table.allpol.(i))) + (lm h))) + && (h.num < i + || not (m = ppcm_mon + (lm (table.allpol.(j))) + (lm h)))) lp let infobuch p q = @@ -668,18 +668,18 @@ let pbuchf table metadata cur_pb homogeneous (lp, lpc) p = infobuch lp lpc; match Heap.pop lpc with | None -> - test_dans_ideal cur_pb table metadata p lp len0 + test_dans_ideal cur_pb table metadata p lp len0 | Some (((i, j), m), lpc2) -> - if critere3 table ((i,j),m) lp lpc2 - then (sinfo "c"; pbuchf cur_pb (lp, lpc2)) - else - let (a0, p0, q0) = spol0 table.allpol.(i) table.allpol.(j) in - if homogeneous && a0 <>[] && deg_hom a0 > deg_hom cur_pb.cur_poly - then (sinfo "h"; pbuchf cur_pb (lp, lpc2)) - else + if critere3 table ((i,j),m) lp lpc2 + then (sinfo "c"; pbuchf cur_pb (lp, lpc2)) + else + let (a0, p0, q0) = spol0 table.allpol.(i) table.allpol.(j) in + if homogeneous && a0 <>[] && deg_hom a0 > deg_hom cur_pb.cur_poly + then (sinfo "h"; pbuchf cur_pb (lp, lpc2)) + else (* let sa = a.sugar in*) match reduce2 table a0 lp with - _, [] -> sinfo "0";pbuchf cur_pb (lp, lpc2) + _, [] -> sinfo "0";pbuchf cur_pb (lp, lpc2) | ca, _ :: _ -> (* info "pair reduced\n";*) let map q = @@ -692,22 +692,22 @@ let pbuchf table metadata cur_pb homogeneous (lp, lpc) p = let (lca, a0) = reduce2_trace table (emultP ca a0) lp lcp in (* info "paire re-reduced";*) let a = new_allpol table a0 in - List.iter2 (fun c q -> coefpoldep_set table a q c) lca lp; - let a0 = a in - info (fun () -> "new polynomial: "^(stringPcut metadata (ppol a0))); + List.iter2 (fun c q -> coefpoldep_set table a q c) lca lp; + let a0 = a in + info (fun () -> "new polynomial: "^(stringPcut metadata (ppol a0))); let nlp = addS a0 lp in - try test_dans_ideal cur_pb table metadata p nlp len0 - with NotInIdealUpdate cur_pb -> - let newlpc = cpairs1 a0 lp lpc2 in - pbuchf cur_pb (nlp, newlpc) + try test_dans_ideal cur_pb table metadata p nlp len0 + with NotInIdealUpdate cur_pb -> + let newlpc = cpairs1 a0 lp lpc2 in + pbuchf cur_pb (nlp, newlpc) in pbuchf cur_pb (lp, lpc) - + let is_homogeneous p = match p with | [] -> true | (a,m)::p1 -> let d = deg m in List.for_all (fun (b,m') -> deg m' =d) p1 - + (* returns c lp = [pn;...;p1] @@ -719,7 +719,7 @@ let is_homogeneous p = lc = [qn+m; ... q1] such that - c*p = sum qi*pi + c*p = sum qi*pi where pn+k = a(n+k,n+k-1)*pn+k-1 + ... + a(n+k,1)* p1 *) diff --git a/plugins/nsatz/nsatz.ml b/plugins/nsatz/nsatz.ml index 71a3132283..9ba83c0843 100644 --- a/plugins/nsatz/nsatz.ml +++ b/plugins/nsatz/nsatz.ml @@ -239,7 +239,7 @@ let rec parse_request lp = match Constr.kind lp with | App (_,[|_|]) -> [] | App (_,[|_;p;lp1|]) -> - (parse_term p)::(parse_request lp1) + (parse_term p)::(parse_request lp1) |_-> assert false let set_nvars_term nvars t = @@ -266,7 +266,7 @@ module Poly = Polynom.Make(Coef) module PIdeal = Ideal.Make(Poly) open PIdeal -(* term to sparse polynomial +(* term to sparse polynomial variables <=np are in the coefficients *) @@ -278,22 +278,22 @@ let term_pol_sparse nvars np t= match t with | Zero -> zeroP | Const r -> - if Num.eq_num r num_0 - then zeroP - else polconst d (Poly.Pint (Coef.of_num r)) + if Num.eq_num r num_0 + then zeroP + else polconst d (Poly.Pint (Coef.of_num r)) | Var v -> - let v = int_of_string v in - if v <= np - then polconst d (Poly.x v) - else gen d v + let v = int_of_string v in + if v <= np + then polconst d (Poly.x v) + else gen d v | Opp t1 -> oppP (aux t1) | Add (t1,t2) -> plusP (aux t1) (aux t2) | Sub (t1,t2) -> plusP (aux t1) (oppP (aux t2)) | Mul (t1,t2) -> multP (aux t1) (aux t2) | Pow (t1,n) -> puisP (aux t1) n - in + in (* info ("donne: "^(stringP res)^"\n");*) - res + res in let res= aux t in res @@ -321,25 +321,25 @@ let pol_sparse_to_term n2 p = let m = Ideal.Monomial.repr m in let n = (Array.length m)-1 in let (i0,e0) = - List.fold_left (fun (r,d) (a,m) -> + List.fold_left (fun (r,d) (a,m) -> let m = Ideal.Monomial.repr m in - let i0= ref 0 in - for k=1 to n do - if m.(k)>0 - then i0:=k - done; - if Int.equal !i0 0 - then (r,d) - else if !i0 > r - then (!i0, m.(!i0)) - else if Int.equal !i0 r && m.(!i0)<d - then (!i0, m.(!i0)) - else (r,d)) - (0,0) - p in + let i0= ref 0 in + for k=1 to n do + if m.(k)>0 + then i0:=k + done; + if Int.equal !i0 0 + then (r,d) + else if !i0 > r + then (!i0, m.(!i0)) + else if Int.equal !i0 r && m.(!i0)<d + then (!i0, m.(!i0)) + else (r,d)) + (0,0) + p in if Int.equal i0 0 then - let mp = polrec_to_term a in + let mp = polrec_to_term a in if List.is_empty p1 then mp else add (mp, aux p1) else let fold (p1, p2) (a, m) = @@ -352,11 +352,11 @@ let pol_sparse_to_term n2 p = (p1, (a, m) :: p2) in let (p1, p2) = List.fold_left fold ([], []) p in - let vm = - if Int.equal e0 1 - then Var (string_of_int (i0)) - else pow (Var (string_of_int (i0)),e0) in - add (mul(vm, aux (List.rev p1)), aux (List.rev p2)) + let vm = + if Int.equal e0 1 + then Var (string_of_int (i0)) + else pow (Var (string_of_int (i0)),e0) in + add (mul(vm, aux (List.rev p1)), aux (List.rev p2)) in (*info "-> pol_sparse_to_term\n";*) aux p @@ -410,34 +410,34 @@ open Ideal *) let clean_pol lp = let t = Hashpol.create 12 in - let find p = try Hashpol.find t p - with + let find p = try Hashpol.find t p + with Not_found -> Hashpol.add t p true; false in let rec aux lp = match lp with | [] -> [], [] - | p :: lp1 -> + | p :: lp1 -> let clp, lb = aux lp1 in - if equal p zeroP || find p then clp, true::lb + if equal p zeroP || find p then clp, true::lb else (p :: clp, false::lb) in aux lp -(* Expand the list of polynomials lp putting zeros where the list of - booleans lb indicates there is a missing element +(* Expand the list of polynomials lp putting zeros where the list of + booleans lb indicates there is a missing element Warning: the expansion is relative to the end of the list in reversed order lp cannot have less elements than lb *) let expand_pol lb lp = - let rec aux lb lp = + let rec aux lb lp = match lb with | [] -> lp | true :: lb1 -> zeroP :: aux lb1 lp | false :: lb1 -> match lp with [] -> assert false - | p :: lp1 -> p :: aux lb1 lp1 + | p :: lp1 -> p :: aux lb1 lp1 in List.rev (aux lb (List.rev lp)) let theoremedeszeros_termes lp = @@ -446,21 +446,21 @@ let theoremedeszeros_termes lp = | Const (Int sugarparam)::Const (Int nparam)::lp -> ((match sugarparam with |0 -> sinfo "computation without sugar"; - lexico:=false; + lexico:=false; |1 -> sinfo "computation with sugar"; - lexico:=false; + lexico:=false; |2 -> sinfo "ordre lexico computation without sugar"; - lexico:=true; + lexico:=true; |3 -> sinfo "ordre lexico computation with sugar"; - lexico:=true; + lexico:=true; |4 -> sinfo "computation without sugar, division by pairs"; - lexico:=false; + lexico:=false; |5 -> sinfo "computation with sugar, division by pairs"; - lexico:=false; + lexico:=false; |6 -> sinfo "ordre lexico computation without sugar, division by pairs"; - lexico:=true; + lexico:=true; |7 -> sinfo "ordre lexico computation with sugar, division by pairs"; - lexico:=true; + lexico:=true; | _ -> user_err Pp.(str "nsatz: bad parameter") ); let lvar = List.init nvars (fun i -> Printf.sprintf "x%i" (i + 1)) in @@ -471,32 +471,32 @@ let theoremedeszeros_termes lp = match lp with | [] -> assert false | p::lp1 -> - let lpol = List.rev lp1 in + let lpol = List.rev lp1 in (* preprocessing : we remove zero polynomials and duplicate that are not handled by in_ideal - lb is kept in order to fix the certificate in the post-processing + lb is kept in order to fix the certificate in the post-processing *) - let lpol, lb = clean_pol lpol in - let cert = theoremedeszeros metadata nvars lpol p in + let lpol, lb = clean_pol lpol in + let cert = theoremedeszeros metadata nvars lpol p in sinfo "cert ok"; - let lc = cert.last_comb::List.rev cert.gb_comb in - match remove_zeros lc with - | [] -> assert false - | (lq::lci) -> + let lc = cert.last_comb::List.rev cert.gb_comb in + match remove_zeros lc with + | [] -> assert false + | (lq::lci) -> (* post-processing : we apply the correction for the last line *) let lq = expand_pol lb lq in - (* lci commence par les nouveaux polynomes *) - let m = nvars in - let c = pol_sparse_to_term m (polconst m cert.coef) in - let r = Pow(Zero,cert.power) in - let lci = List.rev lci in + (* lci commence par les nouveaux polynomes *) + let m = nvars in + let c = pol_sparse_to_term m (polconst m cert.coef) in + let r = Pow(Zero,cert.power) in + let lci = List.rev lci in (* post-processing we apply the correction for the other lines *) - let lci = List.map (expand_pol lb) lci in - let lci = List.map (List.map (pol_sparse_to_term m)) lci in - let lq = List.map (pol_sparse_to_term m) lq in - info (fun () -> Printf.sprintf "number of parameters: %i" nparam); - sinfo "term computed"; - (c,r,lci,lq) + let lci = List.map (expand_pol lb) lci in + let lci = List.map (List.map (pol_sparse_to_term m)) lci in + let lq = List.map (pol_sparse_to_term m) lq in + info (fun () -> Printf.sprintf "number of parameters: %i" nparam); + sinfo "term computed"; + (c,r,lci,lq) ) |_ -> assert false @@ -526,13 +526,13 @@ let nsatz lpol = let res = List.fold_right (fun lt r -> - let ltterm = - List.fold_right - (fun t r -> - mkt_app lcons [mkt_app tpexpr [Lazy.force tz];t;r]) - lt - (mkt_app lnil [mkt_app tpexpr [Lazy.force tz]]) in - mkt_app lcons [tlp ();ltterm;r]) + let ltterm = + List.fold_right + (fun t r -> + mkt_app lcons [mkt_app tpexpr [Lazy.force tz];t;r]) + lt + (mkt_app lnil [mkt_app tpexpr [Lazy.force tz]]) in + mkt_app lcons [tlp ();ltterm;r]) res (mkt_app lnil [tlp ()]) in sinfo "term computed"; diff --git a/plugins/nsatz/polynom.ml b/plugins/nsatz/polynom.ml index 071c74ab9b..9a22f39f84 100644 --- a/plugins/nsatz/polynom.ml +++ b/plugins/nsatz/polynom.ml @@ -181,7 +181,7 @@ let norm p = match p with let d = (Array.length a -1) in let n = ref d in while !n>0 && (equal a.(!n) (Pint coef0)) do - n:=!n-1; + n:=!n-1; done; if !n<0 then Pint coef0 else if Int.equal !n 0 then a.(0) @@ -222,7 +222,7 @@ let coef v i p = let rec plusP p q = let res = (match (p,q) with - (Pint a,Pint b) -> Pint (C.plus a b) + (Pint a,Pint b) -> Pint (C.plus a b) |(Pint a, Prec (y,q1)) -> let q2=Array.map copyP q1 in q2.(0)<- plusP p q1.(0); Prec (y,q2) @@ -317,7 +317,7 @@ let deriv v p = else (let p2 = Array.make d (Pint coef0) in for i=0 to d-1 do - p2.(i)<- multP (Pint (coef_of_int (i+1))) p1.(i+1); + p2.(i)<- multP (Pint (coef_of_int (i+1))) p1.(i+1); done; Prec (x,p2)) | Prec(x,p1)-> Pint coef0 @@ -416,37 +416,37 @@ let rec string_of_Pcut p = for i=(Array.length t)-1 downto 1 do if (!nsP)<0 then (sp:="..."; - if not (!fin) then s:=(!s)^"+"^(!sp); - fin:=true) + if not (!fin) then s:=(!s)^"+"^(!sp); + fin:=true) else ( - let si=string_of_Pcut t.(i) in - sp:=""; - if Int.equal i 1 - then ( - if not (String.equal si "0") - then (nsP:=(!nsP)-1; - if String.equal si "1" - then sp:=v - else - (if (String.contains si '+') - then sp:="("^si^")*"^v - else sp:=si^"*"^v))) - else ( - if not (String.equal si "0") - then (nsP:=(!nsP)-1; - if String.equal si "1" - then sp:=v^"^"^(string_of_int i) - else (if (String.contains si '+') - then sp:="("^si^")*"^v^"^"^(string_of_int i) - else sp:=si^"*"^v^"^"^(string_of_int i)))); - if not (String.is_empty !sp) && not (!fin) - then (nsP:=(!nsP)-1; - if String.is_empty !s - then s:=!sp - else s:=(!s)^"+"^(!sp))); + let si=string_of_Pcut t.(i) in + sp:=""; + if Int.equal i 1 + then ( + if not (String.equal si "0") + then (nsP:=(!nsP)-1; + if String.equal si "1" + then sp:=v + else + (if (String.contains si '+') + then sp:="("^si^")*"^v + else sp:=si^"*"^v))) + else ( + if not (String.equal si "0") + then (nsP:=(!nsP)-1; + if String.equal si "1" + then sp:=v^"^"^(string_of_int i) + else (if (String.contains si '+') + then sp:="("^si^")*"^v^"^"^(string_of_int i) + else sp:=si^"*"^v^"^"^(string_of_int i)))); + if not (String.is_empty !sp) && not (!fin) + then (nsP:=(!nsP)-1; + if String.is_empty !s + then s:=!sp + else s:=(!s)^"+"^(!sp))); done; if String.is_empty !s then (nsP:=(!nsP)-1; - (s:="0")); + (s:="0")); !s let to_string p = @@ -471,10 +471,10 @@ let rec quo_rem_pol p q x = if Int.equal x 0 then (match (p,q) with |(Pint a, Pint b) -> - if C.equal (C.modulo a b) coef0 + if C.equal (C.modulo a b) coef0 then (Pint (C.div a b), cf0) else failwith "div_pol1" - |_ -> assert false) + |_ -> assert false) else let m = deg x q in let b = coefDom x q in @@ -483,14 +483,14 @@ let rec quo_rem_pol p q x = let s = ref cf0 in let continue =ref true in while (!continue) && (not (equal !r cf0)) do - let n = deg x !r in - if n<m - then continue:=false - else ( + let n = deg x !r in + if n<m + then continue:=false + else ( let a = coefDom x !r in let p1 = remP x !r in (* r = a*x^n+p1 *) let c = div_pol a b (x-1) in (* a = c*b *) - let s1 = c @@ ((monome x (n-m))) in + let s1 = c @@ ((monome x (n-m))) in s:= plusP (!s) s1; r:= p1 -- (s1 @@ q1); ) @@ -503,11 +503,11 @@ and div_pol p q x = if equal r cf0 then s else failwith ("div_pol:\n" - ^"p:"^(to_string p)^"\n" - ^"q:"^(to_string q)^"\n" - ^"r:"^(to_string r)^"\n" - ^"x:"^(string_of_int x)^"\n" - ) + ^"p:"^(to_string p)^"\n" + ^"q:"^(to_string q)^"\n" + ^"r:"^(to_string r)^"\n" + ^"x:"^(string_of_int x)^"\n" + ) let divP p q= let x = max (max_var_pol p) (max_var_pol q) in div_pol p q x @@ -534,29 +534,29 @@ let pseudo_div p q x = Pint _ -> (cf0, q,1, p) | Prec (v,q1) when not (Int.equal x v) -> (cf0, q,1, p) | Prec (v,q1) -> - ( - (* pr "pseudo_division: c^d*p = s*q + r";*) - let delta = ref 0 in - let r = ref p in - let c = coefDom x q in - let q1 = remP x q in - let d' = deg x q in - let s = ref cf0 in - while (deg x !r)>=(deg x q) do - let d = deg x !r in - let a = coefDom x !r in - let r1=remP x !r in - let u = a @@ ((monome x (d-d'))) in - r:=(c @@ r1) -- (u @@ q1); - s:=plusP (c @@ (!s)) u; - delta := (!delta) + 1; - done; - (* - pr ("deg d: "^(string_of_int (!delta))^", deg c: "^(string_of_int (deg_total c))); - pr ("deg r:"^(string_of_int (deg_total !r))); - *) - (!r,c,!delta, !s) - ) + ( + (* pr "pseudo_division: c^d*p = s*q + r";*) + let delta = ref 0 in + let r = ref p in + let c = coefDom x q in + let q1 = remP x q in + let d' = deg x q in + let s = ref cf0 in + while (deg x !r)>=(deg x q) do + let d = deg x !r in + let a = coefDom x !r in + let r1=remP x !r in + let u = a @@ ((monome x (d-d'))) in + r:=(c @@ r1) -- (u @@ q1); + s:=plusP (c @@ (!s)) u; + delta := (!delta) + 1; + done; + (* + pr ("deg d: "^(string_of_int (!delta))^", deg c: "^(string_of_int (deg_total c))); + pr ("deg r:"^(string_of_int (deg_total !r))); + *) + (!r,c,!delta, !s) + ) (* gcd with subresultants *) @@ -581,28 +581,28 @@ and pgcd_coef_pol c p x = and pgcd_pol_rec p q x = match (p,q) with - (Pint a,Pint b) -> Pint (C.pgcd (C.abs a) (C.abs b)) + (Pint a,Pint b) -> Pint (C.pgcd (C.abs a) (C.abs b)) |_ -> - if equal p cf0 - then q - else if equal q cf0 - then p - else if Int.equal (deg x q) 0 - then pgcd_coef_pol q p x - else if Int.equal (deg x p) 0 - then pgcd_coef_pol p q x - else ( - let a = content_pol p x in - let b = content_pol q x in - let c = pgcd_pol_rec a b (x-1) in - pr (string_of_int x); - let p1 = div_pol p c x in - let q1 = div_pol q c x in - let r = gcd_sub_res p1 q1 x in - let cr = content_pol r x in - let res = c @@ (div_pol r cr x) in - res - ) + if equal p cf0 + then q + else if equal q cf0 + then p + else if Int.equal (deg x q) 0 + then pgcd_coef_pol q p x + else if Int.equal (deg x p) 0 + then pgcd_coef_pol p q x + else ( + let a = content_pol p x in + let b = content_pol q x in + let c = pgcd_pol_rec a b (x-1) in + pr (string_of_int x); + let p1 = div_pol p c x in + let q1 = div_pol q c x in + let r = gcd_sub_res p1 q1 x in + let cr = content_pol r x in + let res = c @@ (div_pol r cr x) in + res + ) (* Sub-résultants: @@ -630,10 +630,10 @@ and gcd_sub_res p q x = if d<d' then gcd_sub_res q p x else - let delta = d-d' in - let c' = coefDom x q in - let r = snd (quo_rem_pol (((oppP c')^^(delta+1))@@p) (oppP q) x) in - gcd_sub_res_rec q r (c'^^delta) c' d' x + let delta = d-d' in + let c' = coefDom x q in + let r = snd (quo_rem_pol (((oppP c')^^(delta+1))@@p) (oppP q) x) in + gcd_sub_res_rec q r (c'^^delta) c' d' x and gcd_sub_res_rec p q s c d x = if equal q cf0 diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml index 00ea9b6a66..dcd85401d6 100644 --- a/plugins/omega/coq_omega.ml +++ b/plugins/omega/coq_omega.ml @@ -373,11 +373,11 @@ let mk_integer n = let rec loop n = if n =? one then Lazy.force coq_xH else mkApp((if n mod two =? zero then Lazy.force coq_xO else Lazy.force coq_xI), - [| loop (n/two) |]) + [| loop (n/two) |]) in if n =? zero then Lazy.force coq_Z0 else mkApp ((if n >? zero then Lazy.force coq_Zpos else Lazy.force coq_Zneg), - [| loop (abs n) |]) + [| loop (abs n) |]) type omega_constant = | Zplus | Zmult | Zminus | Zsucc | Zopp | Zpred @@ -494,11 +494,11 @@ let context sigma operation path (t : constr) = | (p, Cast (c,k,t)) -> mkCast (loop i p c,k,t) | ([], _) -> operation i t | ((P_APP n :: p), App (f,v)) -> - let v' = Array.copy v in - v'.(pred n) <- loop i p v'.(pred n); mkApp (f, v') + let v' = Array.copy v in + v'.(pred n) <- loop i p v'.(pred n); mkApp (f, v') | (p, Fix ((_,n as ln),(tys,lna,v))) -> - let l = Array.length v in - let v' = Array.copy v in + let l = Array.length v in + let v' = Array.copy v in v'.(n)<- loop (Util.(+) i l) p v.(n); (mkFix (ln,(tys,lna,v'))) | ((P_TYPE :: p), Prod (n,t,c)) -> (mkProd (n,loop i p t,c)) @@ -507,7 +507,7 @@ let context sigma operation path (t : constr) = | ((P_TYPE :: p), LetIn (n,b,t,c)) -> (mkLetIn (n,b,loop i p t,c)) | (p, _) -> - failwith ("abstract_path " ^ string_of_int(List.length p)) + failwith ("abstract_path " ^ string_of_int(List.length p)) in loop 1 path t @@ -521,7 +521,7 @@ let occurrence sigma path (t : constr) = | ((P_TYPE :: p), Lambda (n,term,c)) -> loop p term | ((P_TYPE :: p), LetIn (n,b,term,c)) -> loop p term | (p, _) -> - failwith ("occurrence " ^ string_of_int(List.length p)) + failwith ("occurrence " ^ string_of_int(List.length p)) in loop path t @@ -575,9 +575,9 @@ let compile name kind = let rec loop accu = function | Oplus(Otimes(Oatom v,Oz n),r) -> loop ({v=intern_id v; c=n} :: accu) r | Oz n -> - let id = new_id () in - tag_hypothesis name id; - {kind = kind; body = List.rev accu; constant = n; id = id} + let id = new_id () in + tag_hypothesis name id; + {kind = kind; body = List.rev accu; constant = n; id = id} | _ -> anomaly (Pp.str "compile_equation.") in loop [] @@ -608,9 +608,9 @@ let clever_rewrite_base_poly typ p result theorem = mkArrow typ Sorts.Relevant mkProp, mkLambda (make_annot (Name (Id.of_string "H")) Sorts.Relevant, - applist (mkRel 1,[result]), - mkApp (Lazy.force coq_eq_ind_r, - [| typ; result; mkRel 2; mkRel 1; occ; theorem |]))), + applist (mkRel 1,[result]), + mkApp (Lazy.force coq_eq_ind_r, + [| typ; result; mkRel 2; mkRel 1; occ; theorem |]))), [abstracted]) in let argt = mkApp (abstracted, [|result|]) in @@ -692,51 +692,51 @@ let simpl_coeffs path_init path_k = let rec shuffle p (t1,t2) = match t1,t2 with | Oplus(l1,r1), Oplus(l2,r2) -> - if weight l1 > weight l2 then + if weight l1 > weight l2 then let (tac,t') = shuffle (P_APP 2 :: p) (r1,t2) in (clever_rewrite p [[P_APP 1;P_APP 1]; - [P_APP 1; P_APP 2];[P_APP 2]] + [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: tac, Oplus(l1,t')) - else - let (tac,t') = shuffle (P_APP 2 :: p) (t1,r2) in + else + let (tac,t') = shuffle (P_APP 2 :: p) (t1,r2) in (clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 1];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zplus_permute) :: tac, Oplus(l2,t')) | Oplus(l1,r1), t2 -> - if weight l1 > weight t2 then + if weight l1 > weight t2 then let (tac,t') = shuffle (P_APP 2 :: p) (r1,t2) in clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: tac, Oplus(l1, t') - else + else [clever_rewrite p [[P_APP 1];[P_APP 2]] - (Lazy.force coq_fast_Zplus_comm)], + (Lazy.force coq_fast_Zplus_comm)], Oplus(t2,t1) | t1,Oplus(l2,r2) -> - if weight l2 > weight t1 then + if weight l2 > weight t1 then let (tac,t') = shuffle (P_APP 2 :: p) (t1,r2) in clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 1];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zplus_permute) :: tac, Oplus(l2,t') - else [],Oplus(t1,t2) + else [],Oplus(t1,t2) | Oz t1,Oz t2 -> - [focused_simpl p], Oz(Bigint.add t1 t2) + [focused_simpl p], Oz(Bigint.add t1 t2) | t1,t2 -> - if weight t1 < weight t2 then + if weight t1 < weight t2 then [clever_rewrite p [[P_APP 1];[P_APP 2]] - (Lazy.force coq_fast_Zplus_comm)], - Oplus(t2,t1) - else [],Oplus(t1,t2) + (Lazy.force coq_fast_Zplus_comm)], + Oplus(t2,t1) + else [],Oplus(t1,t2) let shuffle_mult p_init k1 e1 k2 e2 = let rec loop p = function | (({c=c1;v=v1}::l1) as l1'),(({c=c2;v=v2}::l2) as l2') -> - if Int.equal v1 v2 then + if Int.equal v1 v2 then let tac = clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1; P_APP 1]; [P_APP 1; P_APP 1; P_APP 1; P_APP 2]; @@ -746,15 +746,15 @@ let shuffle_mult p_init k1 e1 k2 e2 = [P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA10) - in + in if Bigint.add (Bigint.mult k1 c1) (Bigint.mult k2 c2) =? zero then let tac' = - clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] + clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zred_factor5) in - tac :: focused_simpl (P_APP 2::P_APP 1:: p) :: tac' :: - loop p (l1,l2) + tac :: focused_simpl (P_APP 2::P_APP 1:: p) :: tac' :: + loop p (l1,l2) else tac :: loop (P_APP 2 :: p) (l1,l2) - else if v1 > v2 then + else if v1 > v2 then clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1; P_APP 1]; [P_APP 1; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 1; P_APP 2]; @@ -762,7 +762,7 @@ let shuffle_mult p_init k1 e1 k2 e2 = [P_APP 1; P_APP 2]] (Lazy.force coq_fast_OMEGA11) :: loop (P_APP 2 :: p) (l1,l2') - else + else clever_rewrite p [[P_APP 2; P_APP 1; P_APP 1; P_APP 1]; [P_APP 2; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1]; @@ -793,7 +793,7 @@ let shuffle_mult p_init k1 e1 k2 e2 = let shuffle_mult_right p_init e1 k2 e2 = let rec loop p = function | (({c=c1;v=v1}::l1) as l1'),(({c=c2;v=v2}::l2) as l2') -> - if Int.equal v1 v2 then + if Int.equal v1 v2 then let tac = clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1]; @@ -803,20 +803,20 @@ let shuffle_mult_right p_init e1 k2 e2 = [P_APP 2; P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]] (Lazy.force coq_fast_OMEGA15) - in + in if Bigint.add c1 (Bigint.mult k2 c2) =? zero then let tac' = clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] (Lazy.force coq_fast_Zred_factor5) - in + in tac :: focused_simpl (P_APP 2::P_APP 1:: p) :: tac' :: loop p (l1,l2) else tac :: loop (P_APP 2 :: p) (l1,l2) - else if v1 > v2 then + else if v1 > v2 then clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: loop (P_APP 2 :: p) (l1,l2') - else + else clever_rewrite p [[P_APP 2; P_APP 1; P_APP 1; P_APP 1]; [P_APP 2; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1]; @@ -844,13 +844,13 @@ let rec shuffle_cancel p = function | [] -> [focused_simpl p] | ({c=c1}::l1) -> let tac = - clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1];[P_APP 1; P_APP 2]; + clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1];[P_APP 1; P_APP 2]; [P_APP 2; P_APP 2]; [P_APP 1; P_APP 1; P_APP 2; P_APP 1]] (if c1 >? zero then - (Lazy.force coq_fast_OMEGA13) - else - (Lazy.force coq_fast_OMEGA14)) + (Lazy.force coq_fast_OMEGA13) + else + (Lazy.force coq_fast_OMEGA14)) in tac :: shuffle_cancel p l1 @@ -875,7 +875,7 @@ let scalar_norm p_init = let rec loop p = function | [] -> [simpl_coeffs p_init p] | (_::l) -> - clever_rewrite p + clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1];[P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_OMEGA16) :: loop (P_APP 2 :: p) l @@ -886,9 +886,9 @@ let norm_add p_init = let rec loop p = function | [] -> [simpl_coeffs p_init p] | _:: l -> - clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] + clever_rewrite p [[P_APP 1;P_APP 1]; [P_APP 1; P_APP 2];[P_APP 2]] (Lazy.force coq_fast_Zplus_assoc_reverse) :: - loop (P_APP 2 :: p) l + loop (P_APP 2 :: p) l in loop p_init @@ -896,7 +896,7 @@ let scalar_norm_add p_init = let rec loop p = function | [] -> [simpl_coeffs p_init p] | _ :: l -> - clever_rewrite p + clever_rewrite p [[P_APP 1; P_APP 1; P_APP 1; P_APP 1]; [P_APP 1; P_APP 1; P_APP 1; P_APP 2]; [P_APP 1; P_APP 1; P_APP 2]; [P_APP 2]; [P_APP 1; P_APP 2]] @@ -936,24 +936,24 @@ let rec transform sigma p t = in try match destructurate_term sigma t with | Kapp(Zplus,[t1;t2]) -> - let tac1,t1' = transform sigma (P_APP 1 :: p) t1 - and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in - let tac,t' = shuffle p (t1',t2') in - tac1 @ tac2 @ tac, t' + let tac1,t1' = transform sigma (P_APP 1 :: p) t1 + and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in + let tac,t' = shuffle p (t1',t2') in + tac1 @ tac2 @ tac, t' | Kapp(Zminus,[t1;t2]) -> - let tac,t = - transform sigma p - (mkApp (Lazy.force coq_Zplus, - [| t1; (mkApp (Lazy.force coq_Zopp, [| t2 |])) |])) in - unfold sp_Zminus :: tac,t + let tac,t = + transform sigma p + (mkApp (Lazy.force coq_Zplus, + [| t1; (mkApp (Lazy.force coq_Zopp, [| t2 |])) |])) in + unfold sp_Zminus :: tac,t | Kapp(Zsucc,[t1]) -> - let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, - [| t1; mk_integer one |])) in - unfold sp_Zsucc :: tac,t + let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, + [| t1; mk_integer one |])) in + unfold sp_Zsucc :: tac,t | Kapp(Zpred,[t1]) -> - let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, - [| t1; mk_integer negone |])) in - unfold sp_Zpred :: tac,t + let tac,t = transform sigma p (mkApp (Lazy.force coq_Zplus, + [| t1; mk_integer negone |])) in + unfold sp_Zpred :: tac,t | Kapp(Zmult,[t1;t2]) -> let tac1,t1' = transform sigma (P_APP 1 :: p) t1 and tac2,t2' = transform sigma (P_APP 2 :: p) t2 in @@ -961,8 +961,8 @@ let rec transform sigma p t = | (_,Oz n) -> let tac,t' = scalar p n t1' in tac1 @ tac2 @ tac,t' | (Oz n,_) -> let sym = - clever_rewrite p [[P_APP 1];[P_APP 2]] - (Lazy.force coq_fast_Zmult_comm) in + clever_rewrite p [[P_APP 1];[P_APP 2]] + (Lazy.force coq_fast_Zmult_comm) in let tac,t' = scalar p n t2' in tac1 @ tac2 @ (sym :: tac),t' | _ -> default false t end @@ -981,26 +981,26 @@ let rec transform sigma p t = let shrink_pair p f1 f2 = match f1,f2 with | Oatom v,Oatom _ -> - let r = Otimes(Oatom v,Oz two) in - clever_rewrite p [[P_APP 1]] (Lazy.force coq_fast_Zred_factor1), r + let r = Otimes(Oatom v,Oz two) in + clever_rewrite p [[P_APP 1]] (Lazy.force coq_fast_Zred_factor1), r | Oatom v, Otimes(_,c2) -> - let r = Otimes(Oatom v,Oplus(c2,Oz one)) in - clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 2]] - (Lazy.force coq_fast_Zred_factor2), r + let r = Otimes(Oatom v,Oplus(c2,Oz one)) in + clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 2]] + (Lazy.force coq_fast_Zred_factor2), r | Otimes (v1,c1),Oatom v -> - let r = Otimes(Oatom v,Oplus(c1,Oz one)) in - clever_rewrite p [[P_APP 2];[P_APP 1;P_APP 2]] + let r = Otimes(Oatom v,Oplus(c1,Oz one)) in + clever_rewrite p [[P_APP 2];[P_APP 1;P_APP 2]] (Lazy.force coq_fast_Zred_factor3), r | Otimes (Oatom v,c1),Otimes (v2,c2) -> - let r = Otimes(Oatom v,Oplus(c1,c2)) in - clever_rewrite p + let r = Otimes(Oatom v,Oplus(c1,c2)) in + clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 1;P_APP 2];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zred_factor4),r | t1,t2 -> - begin - oprint t1; print_newline (); oprint t2; print_newline (); + begin + oprint t1; print_newline (); oprint t2; print_newline (); flush stdout; CErrors.user_err Pp.(str "shrink.1") - end + end let reduce_factor p = function | Oatom v -> @@ -1010,8 +1010,8 @@ let reduce_factor p = function | Otimes(Oatom v,c) -> let rec compute = function | Oz n -> n - | Oplus(t1,t2) -> Bigint.add (compute t1) (compute t2) - | _ -> CErrors.user_err Pp.(str "condense.1") + | Oplus(t1,t2) -> Bigint.add (compute t1) (compute t2) + | _ -> CErrors.user_err Pp.(str "condense.1") in [focused_simpl (P_APP 2 :: p)], Otimes(Oatom v,Oz(compute c)) | t -> oprint t; CErrors.user_err Pp.(str "reduce_factor.1") @@ -1019,29 +1019,29 @@ let reduce_factor p = function let rec condense p = function | Oplus(f1,(Oplus(f2,r) as t)) -> if Int.equal (weight f1) (weight f2) then begin - let shrink_tac,t = shrink_pair (P_APP 1 :: p) f1 f2 in - let assoc_tac = + let shrink_tac,t = shrink_pair (P_APP 1 :: p) f1 f2 in + let assoc_tac = clever_rewrite p [[P_APP 1];[P_APP 2;P_APP 1];[P_APP 2;P_APP 2]] (Lazy.force coq_fast_Zplus_assoc) in - let tac_list,t' = condense p (Oplus(t,r)) in - (assoc_tac :: shrink_tac :: tac_list), t' + let tac_list,t' = condense p (Oplus(t,r)) in + (assoc_tac :: shrink_tac :: tac_list), t' end else begin - let tac,f = reduce_factor (P_APP 1 :: p) f1 in - let tac',t' = condense (P_APP 2 :: p) t in - (tac @ tac'), Oplus(f,t') + let tac,f = reduce_factor (P_APP 1 :: p) f1 in + let tac',t' = condense (P_APP 2 :: p) t in + (tac @ tac'), Oplus(f,t') end | Oplus(f1,Oz n) -> let tac,f1' = reduce_factor (P_APP 1 :: p) f1 in tac,Oplus(f1',Oz n) | Oplus(f1,f2) -> if Int.equal (weight f1) (weight f2) then begin - let tac_shrink,t = shrink_pair p f1 f2 in - let tac,t' = condense p t in - tac_shrink :: tac,t' + let tac_shrink,t = shrink_pair p f1 f2 in + let tac,t' = condense p t in + tac_shrink :: tac,t' end else begin - let tac,f = reduce_factor (P_APP 1 :: p) f1 in - let tac',t' = condense (P_APP 2 :: p) f2 in - (tac @ tac'),Oplus(f,t') + let tac,f = reduce_factor (P_APP 1 :: p) f1 in + let tac',t' = condense (P_APP 2 :: p) f2 in + (tac @ tac'),Oplus(f,t') end | Oz _ as t -> [],t | t -> @@ -1053,8 +1053,8 @@ let rec condense p = function let rec clear_zero p = function | Oplus(Otimes(Oatom v,Oz n),r) when n =? zero -> let tac = - clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] - (Lazy.force coq_fast_Zred_factor5) in + clever_rewrite p [[P_APP 1;P_APP 1];[P_APP 2]] + (Lazy.force coq_fast_Zred_factor5) in let tac',t = clear_zero p r in tac :: tac',t | Oplus(f,r) -> @@ -1069,304 +1069,304 @@ let replay_history tactic_normalisation = let rec loop t : unit Proofview.tactic = match t with | HYP e :: l -> - begin - try - tclTHEN - (Id.List.assoc (hyp_of_tag e.id) tactic_normalisation) - (loop l) - with Not_found -> loop l end + begin + try + tclTHEN + (Id.List.assoc (hyp_of_tag e.id) tactic_normalisation) + (loop l) + with Not_found -> loop l end | NEGATE_CONTRADICT (e2,e1,b) :: l -> - let eq1 = decompile e1 - and eq2 = decompile e2 in - let id1 = hyp_of_tag e1.id - and id2 = hyp_of_tag e2.id in - let k = if b then negone else one in - let p_initial = [P_APP 1;P_TYPE] in - let tac= shuffle_mult_right p_initial e1.body k e2.body in - tclTHENLIST [ - generalize_tac - [mkApp (Lazy.force coq_OMEGA17, [| - val_of eq1; - val_of eq2; - mk_integer k; - mkVar id1; mkVar id2 |])]; - mk_then tac; - (intros_using [aux]); - resolve_id aux; + let eq1 = decompile e1 + and eq2 = decompile e2 in + let id1 = hyp_of_tag e1.id + and id2 = hyp_of_tag e2.id in + let k = if b then negone else one in + let p_initial = [P_APP 1;P_TYPE] in + let tac= shuffle_mult_right p_initial e1.body k e2.body in + tclTHENLIST [ + generalize_tac + [mkApp (Lazy.force coq_OMEGA17, [| + val_of eq1; + val_of eq2; + mk_integer k; + mkVar id1; mkVar id2 |])]; + mk_then tac; + (intros_using [aux]); + resolve_id aux; reflexivity ] | CONTRADICTION (e1,e2) :: l -> - let eq1 = decompile e1 - and eq2 = decompile e2 in - let p_initial = [P_APP 2;P_TYPE] in - let tac = shuffle_cancel p_initial e1.body in - let solve_le = + let eq1 = decompile e1 + and eq2 = decompile e2 in + let p_initial = [P_APP 2;P_TYPE] in + let tac = shuffle_cancel p_initial e1.body in + let solve_le = let not_sup_sup = mkApp (Lazy.force coq_eq, - [| - Lazy.force coq_comparison; - Lazy.force coq_Gt; - Lazy.force coq_Gt |]) - in + [| + Lazy.force coq_comparison; + Lazy.force coq_Gt; + Lazy.force coq_Gt |]) + in tclTHENS - (tclTHENLIST [ - unfold sp_Zle; - simpl_in_concl; - intro; - (absurd not_sup_sup) ]) - [ assumption ; reflexivity ] - in - let theorem = + (tclTHENLIST [ + unfold sp_Zle; + simpl_in_concl; + intro; + (absurd not_sup_sup) ]) + [ assumption ; reflexivity ] + in + let theorem = mkApp (Lazy.force coq_OMEGA2, [| - val_of eq1; val_of eq2; - mkVar (hyp_of_tag e1.id); - mkVar (hyp_of_tag e2.id) |]) - in - Proofview.tclTHEN (tclTHEN (generalize_tac [theorem]) (mk_then tac)) solve_le + val_of eq1; val_of eq2; + mkVar (hyp_of_tag e1.id); + mkVar (hyp_of_tag e2.id) |]) + in + Proofview.tclTHEN (tclTHEN (generalize_tac [theorem]) (mk_then tac)) solve_le | DIVIDE_AND_APPROX (e1,e2,k,d) :: l -> - let id = hyp_of_tag e1.id in - let eq1 = val_of(decompile e1) - and eq2 = val_of(decompile e2) in - let kk = mk_integer k - and dd = mk_integer d in - let rhs = mk_plus (mk_times eq2 kk) dd in - let state_eg = mk_eq eq1 rhs in - let tac = scalar_norm_add [P_APP 3] e2.body in - tclTHENS - (cut state_eg) - [ tclTHENS - (tclTHENLIST [ - (intros_using [aux]); - (generalize_tac - [mkApp (Lazy.force coq_OMEGA1, - [| eq1; rhs; mkVar aux; mkVar id |])]); - (clear [aux;id]); - (intros_using [id]); - (cut (mk_gt kk dd)) ]) - [ tclTHENS - (cut (mk_gt kk izero)) - [ tclTHENLIST [ - (intros_using [aux1; aux2]); - (generalize_tac - [mkApp (Lazy.force coq_Zmult_le_approx, - [| kk;eq2;dd;mkVar aux1;mkVar aux2; mkVar id |])]); - (clear [aux1;aux2;id]); - (intros_using [id]); - (loop l) ]; - tclTHENLIST [ - (unfold sp_Zgt); - simpl_in_concl; - reflexivity ] ]; - tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] + let id = hyp_of_tag e1.id in + let eq1 = val_of(decompile e1) + and eq2 = val_of(decompile e2) in + let kk = mk_integer k + and dd = mk_integer d in + let rhs = mk_plus (mk_times eq2 kk) dd in + let state_eg = mk_eq eq1 rhs in + let tac = scalar_norm_add [P_APP 3] e2.body in + tclTHENS + (cut state_eg) + [ tclTHENS + (tclTHENLIST [ + (intros_using [aux]); + (generalize_tac + [mkApp (Lazy.force coq_OMEGA1, + [| eq1; rhs; mkVar aux; mkVar id |])]); + (clear [aux;id]); + (intros_using [id]); + (cut (mk_gt kk dd)) ]) + [ tclTHENS + (cut (mk_gt kk izero)) + [ tclTHENLIST [ + (intros_using [aux1; aux2]); + (generalize_tac + [mkApp (Lazy.force coq_Zmult_le_approx, + [| kk;eq2;dd;mkVar aux1;mkVar aux2; mkVar id |])]); + (clear [aux1;aux2;id]); + (intros_using [id]); + (loop l) ]; + tclTHENLIST [ + (unfold sp_Zgt); + simpl_in_concl; + reflexivity ] ]; + tclTHENLIST [ unfold sp_Zgt; simpl_in_concl; reflexivity ] ]; - tclTHEN (mk_then tac) reflexivity ] + tclTHEN (mk_then tac) reflexivity ] | NOT_EXACT_DIVIDE (e1,k) :: l -> - let c = floor_div e1.constant k in - let d = Bigint.sub e1.constant (Bigint.mult c k) in - let e2 = {id=e1.id; kind=EQUA;constant = c; + let c = floor_div e1.constant k in + let d = Bigint.sub e1.constant (Bigint.mult c k) in + let e2 = {id=e1.id; kind=EQUA;constant = c; body = map_eq_linear (fun c -> c / k) e1.body } in - let eq2 = val_of(decompile e2) in - let kk = mk_integer k - and dd = mk_integer d in - let tac = scalar_norm_add [P_APP 2] e2.body in - tclTHENS - (cut (mk_gt dd izero)) - [ tclTHENS (cut (mk_gt kk dd)) - [tclTHENLIST [ - (intros_using [aux2;aux1]); - (generalize_tac - [mkApp (Lazy.force coq_OMEGA4, + let eq2 = val_of(decompile e2) in + let kk = mk_integer k + and dd = mk_integer d in + let tac = scalar_norm_add [P_APP 2] e2.body in + tclTHENS + (cut (mk_gt dd izero)) + [ tclTHENS (cut (mk_gt kk dd)) + [tclTHENLIST [ + (intros_using [aux2;aux1]); + (generalize_tac + [mkApp (Lazy.force coq_OMEGA4, [| dd;kk;eq2;mkVar aux1; mkVar aux2 |])]); - (clear [aux1;aux2]); - unfold sp_not; - (intros_using [aux]); - resolve_id aux; - mk_then tac; - assumption ] ; - tclTHENLIST [ - unfold sp_Zgt; - simpl_in_concl; - reflexivity ] ]; + (clear [aux1;aux2]); + unfold sp_not; + (intros_using [aux]); + resolve_id aux; + mk_then tac; + assumption ] ; + tclTHENLIST [ + unfold sp_Zgt; + simpl_in_concl; + reflexivity ] ]; tclTHENLIST [ - unfold sp_Zgt; + unfold sp_Zgt; simpl_in_concl; - reflexivity ] ] + reflexivity ] ] | EXACT_DIVIDE (e1,k) :: l -> - let id = hyp_of_tag e1.id in - let e2 = map_eq_afine (fun c -> c / k) e1 in - let eq1 = val_of(decompile e1) - and eq2 = val_of(decompile e2) in - let kk = mk_integer k in - let state_eq = mk_eq eq1 (mk_times eq2 kk) in - if e1.kind == DISE then + let id = hyp_of_tag e1.id in + let e2 = map_eq_afine (fun c -> c / k) e1 in + let eq1 = val_of(decompile e1) + and eq2 = val_of(decompile e2) in + let kk = mk_integer k in + let state_eq = mk_eq eq1 (mk_times eq2 kk) in + if e1.kind == DISE then let tac = scalar_norm [P_APP 3] e2.body in tclTHENS - (cut state_eq) - [tclTHENLIST [ - (intros_using [aux1]); - (generalize_tac - [mkApp (Lazy.force coq_OMEGA18, + (cut state_eq) + [tclTHENLIST [ + (intros_using [aux1]); + (generalize_tac + [mkApp (Lazy.force coq_OMEGA18, [| eq1;eq2;kk;mkVar aux1; mkVar id |])]); - (clear [aux1;id]); - (intros_using [id]); - (loop l) ]; - tclTHEN (mk_then tac) reflexivity ] - else + (clear [aux1;id]); + (intros_using [id]); + (loop l) ]; + tclTHEN (mk_then tac) reflexivity ] + else let tac = scalar_norm [P_APP 3] e2.body in tclTHENS (cut state_eq) - [ - tclTHENS - (cut (mk_gt kk izero)) - [tclTHENLIST [ - (intros_using [aux2;aux1]); - (generalize_tac - [mkApp (Lazy.force coq_OMEGA3, - [| eq1; eq2; kk; mkVar aux2; mkVar aux1;mkVar id|])]); - (clear [aux1;aux2;id]); - (intros_using [id]); - (loop l) ]; - tclTHENLIST [ - unfold sp_Zgt; + [ + tclTHENS + (cut (mk_gt kk izero)) + [tclTHENLIST [ + (intros_using [aux2;aux1]); + (generalize_tac + [mkApp (Lazy.force coq_OMEGA3, + [| eq1; eq2; kk; mkVar aux2; mkVar aux1;mkVar id|])]); + (clear [aux1;aux2;id]); + (intros_using [id]); + (loop l) ]; + tclTHENLIST [ + unfold sp_Zgt; simpl_in_concl; - reflexivity ] ]; - tclTHEN (mk_then tac) reflexivity ] + reflexivity ] ]; + tclTHEN (mk_then tac) reflexivity ] | (MERGE_EQ(e3,e1,e2)) :: l -> - let id = new_identifier () in - tag_hypothesis id e3; - let id1 = hyp_of_tag e1.id - and id2 = hyp_of_tag e2 in - let eq1 = val_of(decompile e1) - and eq2 = val_of (decompile (negate_eq e1)) in - let tac = - clever_rewrite [P_APP 3] [[P_APP 1]] - (Lazy.force coq_fast_Zopp_eq_mult_neg_1) :: + let id = new_identifier () in + tag_hypothesis id e3; + let id1 = hyp_of_tag e1.id + and id2 = hyp_of_tag e2 in + let eq1 = val_of(decompile e1) + and eq2 = val_of (decompile (negate_eq e1)) in + let tac = + clever_rewrite [P_APP 3] [[P_APP 1]] + (Lazy.force coq_fast_Zopp_eq_mult_neg_1) :: scalar_norm [P_APP 3] e1.body - in - tclTHENS - (cut (mk_eq eq1 (mk_inv eq2))) - [tclTHENLIST [ - (intros_using [aux]); - (generalize_tac [mkApp (Lazy.force coq_OMEGA8, - [| eq1;eq2;mkVar id1;mkVar id2; mkVar aux|])]); - (clear [id1;id2;aux]); - (intros_using [id]); - (loop l) ]; + in + tclTHENS + (cut (mk_eq eq1 (mk_inv eq2))) + [tclTHENLIST [ + (intros_using [aux]); + (generalize_tac [mkApp (Lazy.force coq_OMEGA8, + [| eq1;eq2;mkVar id1;mkVar id2; mkVar aux|])]); + (clear [id1;id2;aux]); + (intros_using [id]); + (loop l) ]; tclTHEN (mk_then tac) reflexivity] | STATE {st_new_eq=e;st_def=def;st_orig=orig;st_coef=m;st_var=v} :: l -> - let id = new_identifier () - and id2 = hyp_of_tag orig.id in - tag_hypothesis id e.id; - let eq1 = val_of(decompile def) - and eq2 = val_of(decompile orig) in - let vid = unintern_id v in - let theorem = + let id = new_identifier () + and id2 = hyp_of_tag orig.id in + tag_hypothesis id e.id; + let eq1 = val_of(decompile def) + and eq2 = val_of(decompile orig) in + let vid = unintern_id v in + let theorem = mkApp (Lazy.force coq_ex, [| - Lazy.force coq_Z; - mkLambda + Lazy.force coq_Z; + mkLambda (make_annot (Name vid) Sorts.Relevant, - Lazy.force coq_Z, - mk_eq (mkRel 1) eq1) |]) - in - let mm = mk_integer m in - let p_initial = [P_APP 2;P_TYPE] in - let tac = + Lazy.force coq_Z, + mk_eq (mkRel 1) eq1) |]) + in + let mm = mk_integer m in + let p_initial = [P_APP 2;P_TYPE] in + let tac = clever_rewrite (P_APP 1 :: P_APP 1 :: P_APP 2 :: p_initial) [[P_APP 1]] (Lazy.force coq_fast_Zopp_eq_mult_neg_1) :: shuffle_mult_right p_initial orig.body m ({c= negone;v= v}::def.body) in - tclTHENS - (cut theorem) - [tclTHENLIST [ - (intros_using [aux]); - (elim_id aux); - (clear [aux]); - (intros_using [vid; aux]); - (generalize_tac - [mkApp (Lazy.force coq_OMEGA9, - [| mkVar vid;eq2;eq1;mm; mkVar id2;mkVar aux |])]); - mk_then tac; - (clear [aux]); - (intros_using [id]); - (loop l) ]; + tclTHENS + (cut theorem) + [tclTHENLIST [ + (intros_using [aux]); + (elim_id aux); + (clear [aux]); + (intros_using [vid; aux]); + (generalize_tac + [mkApp (Lazy.force coq_OMEGA9, + [| mkVar vid;eq2;eq1;mm; mkVar id2;mkVar aux |])]); + mk_then tac; + (clear [aux]); + (intros_using [id]); + (loop l) ]; tclTHEN (exists_tac eq1) reflexivity ] | SPLIT_INEQ(e,(e1,act1),(e2,act2)) :: l -> - let id1 = new_identifier () - and id2 = new_identifier () in - tag_hypothesis id1 e1; tag_hypothesis id2 e2; - let id = hyp_of_tag e.id in - let tac1 = norm_add [P_APP 2;P_TYPE] e.body in - let tac2 = scalar_norm_add [P_APP 2;P_TYPE] e.body in - let eq = val_of(decompile e) in - tclTHENS - (simplest_elim (applist (Lazy.force coq_OMEGA19, [eq; mkVar id]))) - [tclTHENLIST [ mk_then tac1; (intros_using [id1]); (loop act1) ]; + let id1 = new_identifier () + and id2 = new_identifier () in + tag_hypothesis id1 e1; tag_hypothesis id2 e2; + let id = hyp_of_tag e.id in + let tac1 = norm_add [P_APP 2;P_TYPE] e.body in + let tac2 = scalar_norm_add [P_APP 2;P_TYPE] e.body in + let eq = val_of(decompile e) in + tclTHENS + (simplest_elim (applist (Lazy.force coq_OMEGA19, [eq; mkVar id]))) + [tclTHENLIST [ mk_then tac1; (intros_using [id1]); (loop act1) ]; tclTHENLIST [ mk_then tac2; (intros_using [id2]); (loop act2) ]] | SUM(e3,(k1,e1),(k2,e2)) :: l -> - let id = new_identifier () in - tag_hypothesis id e3; - let id1 = hyp_of_tag e1.id - and id2 = hyp_of_tag e2.id in - let eq1 = val_of(decompile e1) - and eq2 = val_of(decompile e2) in - if k1 =? one && e2.kind == EQUA then + let id = new_identifier () in + tag_hypothesis id e3; + let id1 = hyp_of_tag e1.id + and id2 = hyp_of_tag e2.id in + let eq1 = val_of(decompile e1) + and eq2 = val_of(decompile e2) in + if k1 =? one && e2.kind == EQUA then let tac_thm = match e1.kind with - | EQUA -> Lazy.force coq_OMEGA5 - | INEQ -> Lazy.force coq_OMEGA6 - | DISE -> Lazy.force coq_OMEGA20 - in + | EQUA -> Lazy.force coq_OMEGA5 + | INEQ -> Lazy.force coq_OMEGA6 + | DISE -> Lazy.force coq_OMEGA20 + in let kk = mk_integer k2 in let p_initial = if e1.kind == DISE then [P_APP 1; P_TYPE] else [P_APP 2; P_TYPE] in let tac = shuffle_mult_right p_initial e1.body k2 e2.body in tclTHENLIST [ - (generalize_tac + (generalize_tac [mkApp (tac_thm, [| eq1; eq2; kk; mkVar id1; mkVar id2 |])]); - mk_then tac; - (intros_using [id]); - (loop l) + mk_then tac; + (intros_using [id]); + (loop l) ] - else + else let kk1 = mk_integer k1 - and kk2 = mk_integer k2 in - let p_initial = [P_APP 2;P_TYPE] in - let tac= shuffle_mult p_initial k1 e1.body k2 e2.body in + and kk2 = mk_integer k2 in + let p_initial = [P_APP 2;P_TYPE] in + let tac= shuffle_mult p_initial k1 e1.body k2 e2.body in tclTHENS (cut (mk_gt kk1 izero)) - [tclTHENS - (cut (mk_gt kk2 izero)) - [tclTHENLIST [ - (intros_using [aux2;aux1]); - (generalize_tac - [mkApp (Lazy.force coq_OMEGA7, [| - eq1;eq2;kk1;kk2; - mkVar aux1;mkVar aux2; - mkVar id1;mkVar id2 |])]); - (clear [aux1;aux2]); - mk_then tac; - (intros_using [id]); - (loop l) ]; - tclTHENLIST [ - unfold sp_Zgt; + [tclTHENS + (cut (mk_gt kk2 izero)) + [tclTHENLIST [ + (intros_using [aux2;aux1]); + (generalize_tac + [mkApp (Lazy.force coq_OMEGA7, [| + eq1;eq2;kk1;kk2; + mkVar aux1;mkVar aux2; + mkVar id1;mkVar id2 |])]); + (clear [aux1;aux2]); + mk_then tac; + (intros_using [id]); + (loop l) ]; + tclTHENLIST [ + unfold sp_Zgt; simpl_in_concl; - reflexivity ] ]; - tclTHENLIST [ - unfold sp_Zgt; + reflexivity ] ]; + tclTHENLIST [ + unfold sp_Zgt; simpl_in_concl; - reflexivity ] ] + reflexivity ] ] | CONSTANT_NOT_NUL(e,k) :: l -> - tclTHEN ((generalize_tac [mkVar (hyp_of_tag e)])) Equality.discrConcl + tclTHEN ((generalize_tac [mkVar (hyp_of_tag e)])) Equality.discrConcl | CONSTANT_NUL(e) :: l -> - tclTHEN (resolve_id (hyp_of_tag e)) reflexivity + tclTHEN (resolve_id (hyp_of_tag e)) reflexivity | CONSTANT_NEG(e,k) :: l -> - tclTHENLIST [ - (generalize_tac [mkVar (hyp_of_tag e)]); + tclTHENLIST [ + (generalize_tac [mkVar (hyp_of_tag e)]); unfold sp_Zle; - simpl_in_concl; - unfold sp_not; - (intros_using [aux]); - resolve_id aux; - reflexivity + simpl_in_concl; + unfold sp_not; + (intros_using [aux]); + resolve_id aux; + reflexivity ] | _ -> Proofview.tclUNIT () in @@ -1401,29 +1401,29 @@ let destructure_omega env sigma tac_def (id,c) = try match destructurate_prop sigma c with | Kapp(Eq,[typ;t1;t2]) when begin match destructurate_type env sigma typ with Kapp(Z,[]) -> true | _ -> false end -> - let t = mk_plus t1 (mk_inv t2) in - normalize_equation sigma - id EQUA (Lazy.force coq_Zegal_left) 2 t t1 t2 tac_def + let t = mk_plus t1 (mk_inv t2) in + normalize_equation sigma + id EQUA (Lazy.force coq_Zegal_left) 2 t t1 t2 tac_def | Kapp(Zne,[t1;t2]) -> - let t = mk_plus t1 (mk_inv t2) in - normalize_equation sigma - id DISE (Lazy.force coq_Zne_left) 1 t t1 t2 tac_def + let t = mk_plus t1 (mk_inv t2) in + normalize_equation sigma + id DISE (Lazy.force coq_Zne_left) 1 t t1 t2 tac_def | Kapp(Zle,[t1;t2]) -> - let t = mk_plus t2 (mk_inv t1) in - normalize_equation sigma - id INEQ (Lazy.force coq_Zle_left) 2 t t1 t2 tac_def + let t = mk_plus t2 (mk_inv t1) in + normalize_equation sigma + id INEQ (Lazy.force coq_Zle_left) 2 t t1 t2 tac_def | Kapp(Zlt,[t1;t2]) -> - let t = mk_plus (mk_plus t2 (mk_integer negone)) (mk_inv t1) in - normalize_equation sigma - id INEQ (Lazy.force coq_Zlt_left) 2 t t1 t2 tac_def + let t = mk_plus (mk_plus t2 (mk_integer negone)) (mk_inv t1) in + normalize_equation sigma + id INEQ (Lazy.force coq_Zlt_left) 2 t t1 t2 tac_def | Kapp(Zge,[t1;t2]) -> - let t = mk_plus t1 (mk_inv t2) in - normalize_equation sigma - id INEQ (Lazy.force coq_Zge_left) 2 t t1 t2 tac_def + let t = mk_plus t1 (mk_inv t2) in + normalize_equation sigma + id INEQ (Lazy.force coq_Zge_left) 2 t t1 t2 tac_def | Kapp(Zgt,[t1;t2]) -> - let t = mk_plus (mk_plus t1 (mk_integer negone)) (mk_inv t2) in - normalize_equation sigma - id INEQ (Lazy.force coq_Zgt_left) 2 t t1 t2 tac_def + let t = mk_plus (mk_plus t1 (mk_integer negone)) (mk_inv t2) in + normalize_equation sigma + id INEQ (Lazy.force coq_Zgt_left) 2 t t1 t2 tac_def | _ -> tac_def with e when catchable_exception e -> tac_def @@ -1444,25 +1444,25 @@ let coq_omega = let prelude,sys = List.fold_left (fun (tac,sys) (t,(v,th,b)) -> - if b then + if b then let id = new_identifier () in let i = new_id () in tag_hypothesis id i; (tclTHENLIST [ - (simplest_elim (applist (Lazy.force coq_intro_Z, [t]))); - (intros_using [v; id]); - (elim_id id); - (clear [id]); - (intros_using [th;id]); - tac ]), + (simplest_elim (applist (Lazy.force coq_intro_Z, [t]))); + (intros_using [v; id]); + (elim_id id); + (clear [id]); + (intros_using [th;id]); + tac ]), {kind = INEQ; - body = [{v=intern_id v; c=one}]; + body = [{v=intern_id v; c=one}]; constant = zero; id = i} :: sys - else + else (tclTHENLIST [ - (simplest_elim (applist (Lazy.force coq_new_var, [t]))); - (intros_using [v;th]); - tac ]), + (simplest_elim (applist (Lazy.force coq_new_var, [t]))); + (intros_using [v;th]); + tac ]), sys) (Proofview.tclUNIT (),[]) (dump_tables ()) in @@ -1495,61 +1495,61 @@ let nat_inject = try match destructurate_term sigma t with | Kapp(Plus,[t1;t2]) -> tclTHENLIST [ - (clever_rewrite_gen p (mk_plus (mk_inj t1) (mk_inj t2)) + (clever_rewrite_gen p (mk_plus (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_plus),[t1;t2])); - (explore (P_APP 1 :: p) t1); - (explore (P_APP 2 :: p) t2) + (explore (P_APP 1 :: p) t1); + (explore (P_APP 2 :: p) t2) ] | Kapp(Mult,[t1;t2]) -> tclTHENLIST [ - (clever_rewrite_gen p (mk_times (mk_inj t1) (mk_inj t2)) + (clever_rewrite_gen p (mk_times (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_mult),[t1;t2])); - (explore (P_APP 1 :: p) t1); - (explore (P_APP 2 :: p) t2) + (explore (P_APP 1 :: p) t1); + (explore (P_APP 2 :: p) t2) ] | Kapp(Minus,[t1;t2]) -> let id = new_identifier () in tclTHENS (tclTHEN - (simplest_elim (applist (Lazy.force coq_le_gt_dec, [t2;t1]))) - (intros_using [id])) - [ - tclTHENLIST [ - (clever_rewrite_gen p - (mk_minus (mk_inj t1) (mk_inj t2)) + (simplest_elim (applist (Lazy.force coq_le_gt_dec, [t2;t1]))) + (intros_using [id])) + [ + tclTHENLIST [ + (clever_rewrite_gen p + (mk_minus (mk_inj t1) (mk_inj t2)) ((Lazy.force coq_inj_minus1),[t1;t2;mkVar id])); - (loop [id,mkApp (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 zero) + (loop [id,mkApp (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 zero) ((Lazy.force coq_inj_minus2),[t1;t2;mkVar id])) - (loop [id,mkApp (Lazy.force coq_gt, [| t2;t1 |])])) - ] + (loop [id,mkApp (Lazy.force coq_gt, [| t2;t1 |])])) + ] | Kapp(S,[t']) -> let rec is_number t = try match destructurate_term sigma t with - Kapp(S,[t]) -> is_number t + Kapp(S,[t]) -> is_number t | Kapp(O,[]) -> true | _ -> false with e when catchable_exception e -> false - in + in let rec loop p t : unit Proofview.tactic = try match destructurate_term sigma t with - Kapp(S,[t]) -> + Kapp(S,[t]) -> (tclTHEN (clever_rewrite_gen p - (mkApp (Lazy.force coq_Zsucc, [| mk_inj t |])) - ((Lazy.force coq_inj_S),[t])) - (loop (P_APP 1 :: p) t)) + (mkApp (Lazy.force coq_Zsucc, [| mk_inj t |])) + ((Lazy.force coq_inj_S),[t])) + (loop (P_APP 1 :: p) t)) | _ -> explore p t with e when catchable_exception e -> explore p t - in + in if is_number t' then focused_simpl p else loop p t | Kapp(Pred,[t]) -> let t_minus_one = - mkApp (Lazy.force coq_minus, [| t; - mkApp (Lazy.force coq_S, [| Lazy.force coq_O |]) |]) in + mkApp (Lazy.force coq_minus, [| t; + mkApp (Lazy.force coq_S, [| Lazy.force coq_O |]) |]) in tclTHEN (clever_rewrite_gen_nat (P_APP 1 :: p) t_minus_one ((Lazy.force coq_pred_of_minus),[t])) @@ -1562,65 +1562,65 @@ let nat_inject = | [] -> Proofview.tclUNIT () | (i,t)::lit -> Proofview.tclEVARMAP >>= fun sigma -> - begin try match destructurate_prop sigma t with + begin try match destructurate_prop sigma t with Kapp(Le,[t1;t2]) -> - tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_inj_le, [| t1;t2;mkVar i |]) ]); - (explore [P_APP 1; P_TYPE] t1); - (explore [P_APP 2; P_TYPE] t2); - (reintroduce i); - (loop lit) + tclTHENLIST [ + (generalize_tac + [mkApp (Lazy.force coq_inj_le, [| t1;t2;mkVar i |]) ]); + (explore [P_APP 1; P_TYPE] t1); + (explore [P_APP 2; P_TYPE] t2); + (reintroduce i); + (loop lit) ] | Kapp(Lt,[t1;t2]) -> - tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_inj_lt, [| t1;t2;mkVar i |]) ]); - (explore [P_APP 1; P_TYPE] t1); - (explore [P_APP 2; P_TYPE] t2); - (reintroduce i); - (loop lit) + tclTHENLIST [ + (generalize_tac + [mkApp (Lazy.force coq_inj_lt, [| t1;t2;mkVar i |]) ]); + (explore [P_APP 1; P_TYPE] t1); + (explore [P_APP 2; P_TYPE] t2); + (reintroduce i); + (loop lit) ] | Kapp(Ge,[t1;t2]) -> - tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_inj_ge, [| t1;t2;mkVar i |]) ]); - (explore [P_APP 1; P_TYPE] t1); - (explore [P_APP 2; P_TYPE] t2); - (reintroduce i); - (loop lit) + tclTHENLIST [ + (generalize_tac + [mkApp (Lazy.force coq_inj_ge, [| t1;t2;mkVar i |]) ]); + (explore [P_APP 1; P_TYPE] t1); + (explore [P_APP 2; P_TYPE] t2); + (reintroduce i); + (loop lit) ] | Kapp(Gt,[t1;t2]) -> - tclTHENLIST [ - (generalize_tac + tclTHENLIST [ + (generalize_tac [mkApp (Lazy.force coq_inj_gt, [| t1;t2;mkVar i |]) ]); - (explore [P_APP 1; P_TYPE] t1); - (explore [P_APP 2; P_TYPE] t2); - (reintroduce i); - (loop lit) + (explore [P_APP 1; P_TYPE] t1); + (explore [P_APP 2; P_TYPE] t2); + (reintroduce i); + (loop lit) ] | Kapp(Neq,[t1;t2]) -> - tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_inj_neq, [| t1;t2;mkVar i |]) ]); - (explore [P_APP 1; P_TYPE] t1); - (explore [P_APP 2; P_TYPE] t2); - (reintroduce i); - (loop lit) + tclTHENLIST [ + (generalize_tac + [mkApp (Lazy.force coq_inj_neq, [| t1;t2;mkVar i |]) ]); + (explore [P_APP 1; P_TYPE] t1); + (explore [P_APP 2; P_TYPE] t2); + (reintroduce i); + (loop lit) ] | Kapp(Eq,[typ;t1;t2]) -> - if is_conv typ (Lazy.force coq_nat) then + if is_conv typ (Lazy.force coq_nat) then tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_inj_eq, [| t1;t2;mkVar i |]) ]); - (explore [P_APP 2; P_TYPE] t1); - (explore [P_APP 3; P_TYPE] t2); - (reintroduce i); - (loop lit) + (generalize_tac + [mkApp (Lazy.force coq_inj_eq, [| t1;t2;mkVar i |]) ]); + (explore [P_APP 2; P_TYPE] t1); + (explore [P_APP 3; P_TYPE] t2); + (reintroduce i); + (loop lit) ] - else loop lit - | _ -> loop lit - with e when catchable_exception e -> loop lit end + else loop lit + | _ -> loop lit + with e when catchable_exception e -> loop lit end in let hyps_types = Tacmach.New.pf_hyps_types gl in loop (List.rev hyps_types) @@ -1661,17 +1661,17 @@ exception Undecidable let rec decidability env sigma t = match destructurate_prop sigma t with | Kapp(Or,[t1;t2]) -> - mkApp (Lazy.force coq_dec_or, [| t1; t2; + mkApp (Lazy.force coq_dec_or, [| t1; t2; decidability env sigma t1; decidability env sigma t2 |]) | Kapp(And,[t1;t2]) -> - mkApp (Lazy.force coq_dec_and, [| t1; t2; + mkApp (Lazy.force coq_dec_and, [| t1; t2; decidability env sigma t1; decidability env sigma t2 |]) | Kapp(Iff,[t1;t2]) -> - mkApp (Lazy.force coq_dec_iff, [| t1; t2; + mkApp (Lazy.force coq_dec_iff, [| t1; t2; decidability env sigma t1; decidability env sigma t2 |]) | Kimp(t1,t2) -> (* This is the only situation where it's not obvious that [t] - is in Prop. The recursive call on [t2] will ensure that. *) + is in Prop. The recursive call on [t2] will ensure that. *) mkApp (Lazy.force coq_dec_imp, [| t1; t2; decidability env sigma t1; decidability env sigma t2 |]) | Kapp(Not,[t1]) -> @@ -1681,10 +1681,10 @@ let rec decidability env sigma t = | Kapp(Z,[]) -> mkApp (Lazy.force coq_dec_eq, [| t1;t2 |]) | Kapp(Nat,[]) -> mkApp (Lazy.force coq_dec_eq_nat, [| t1;t2 |]) | _ -> raise Undecidable - end + end | Kapp(op,[t1;t2]) -> (try mkApp (Lazy.force (dec_binop op), [| t1; t2 |]) - with Not_found -> raise Undecidable) + with Not_found -> raise Undecidable) | Kapp(False,[]) -> Lazy.force coq_dec_False | Kapp(True,[]) -> Lazy.force coq_dec_True | _ -> raise Undecidable @@ -1736,8 +1736,8 @@ let destructure_hyps = | decl :: lit -> (* variable without body (or !letin_flag isn't set) *) let i = NamedDecl.get_id decl in Proofview.tclEVARMAP >>= fun sigma -> - begin try match destructurate_prop sigma (NamedDecl.get_type decl) with - | Kapp(False,[]) -> elim_id i + begin try match destructurate_prop sigma (NamedDecl.get_type decl) with + | Kapp(False,[]) -> elim_id i | Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit | Kapp(Or,[t1;t2]) -> (tclTHENS @@ -1746,125 +1746,125 @@ let destructure_hyps = onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,t2)::lit))) ]) | Kapp(And,[t1;t2]) -> tclTHEN - (elim_id i) - (onClearedName2 i (fun i1 i2 -> + (elim_id i) + (onClearedName2 i (fun i1 i2 -> loop (LocalAssum (make_annot i1 Sorts.Relevant,t1) :: LocalAssum (make_annot i2 Sorts.Relevant,t2) :: lit))) | Kapp(Iff,[t1;t2]) -> - tclTHEN - (elim_id i) - (onClearedName2 i (fun i1 i2 -> + tclTHEN + (elim_id i) + (onClearedName2 i (fun i1 i2 -> loop (LocalAssum (make_annot i1 Sorts.Relevant,mkArrow t1 Sorts.Relevant t2) :: LocalAssum (make_annot i2 Sorts.Relevant,mkArrow t2 Sorts.Relevant t1) :: lit))) | Kimp(t1,t2) -> - (* t1 and t2 might be in Type rather than Prop. - For t1, the decidability check will ensure being Prop. *) + (* t1 and t2 might be in Type rather than Prop. + For t1, the decidability check will ensure being Prop. *) if Termops.is_Prop sigma (type_of t2) then - let d1 = decidability t1 in - tclTHENLIST [ - (generalize_tac [mkApp (Lazy.force coq_imp_simp, + let d1 = decidability t1 in + tclTHENLIST [ + (generalize_tac [mkApp (Lazy.force coq_imp_simp, [| t1; t2; d1; mkVar i|])]); - (onClearedName i (fun i -> + (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_not t1) t2) :: lit)))) ] else - loop lit + loop lit | Kapp(Not,[t]) -> begin match destructurate_prop sigma t with - Kapp(Or,[t1;t2]) -> + Kapp(Or,[t1;t2]) -> tclTHENLIST [ - (generalize_tac + (generalize_tac [mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]); - (onClearedName i (fun i -> + (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_and (mk_not t1) (mk_not t2)) :: lit)))) ] - | Kapp(And,[t1;t2]) -> - let d1 = decidability t1 in + | Kapp(And,[t1;t2]) -> + let d1 = decidability t1 in tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_not_and, - [| t1; t2; d1; mkVar i |])]); - (onClearedName i (fun i -> + (generalize_tac + [mkApp (Lazy.force coq_not_and, + [| t1; t2; d1; mkVar i |])]); + (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_not t1) (mk_not t2)) :: lit)))) ] - | Kapp(Iff,[t1;t2]) -> - let d1 = decidability t1 in - let d2 = decidability t2 in + | Kapp(Iff,[t1;t2]) -> + let d1 = decidability t1 in + let d2 = decidability t2 in tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_not_iff, - [| t1; t2; d1; d2; mkVar i |])]); - (onClearedName i (fun i -> + (generalize_tac + [mkApp (Lazy.force coq_not_iff, + [| t1; t2; d1; d2; mkVar i |])]); + (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_or (mk_and t1 (mk_not t2)) - (mk_and (mk_not t1) t2)) :: lit)))) + (mk_and (mk_not t1) t2)) :: lit)))) ] - | Kimp(t1,t2) -> - (* t2 must be in Prop otherwise ~(t1->t2) wouldn't be ok. - For t1, being decidable implies being Prop. *) - let d1 = decidability t1 in + | Kimp(t1,t2) -> + (* t2 must be in Prop otherwise ~(t1->t2) wouldn't be ok. + For t1, being decidable implies being Prop. *) + let d1 = decidability t1 in tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_not_imp, - [| t1; t2; d1; mkVar i |])]); - (onClearedName i (fun i -> + (generalize_tac + [mkApp (Lazy.force coq_not_imp, + [| t1; t2; d1; mkVar i |])]); + (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,mk_and t1 (mk_not t2)) :: lit)))) ] - | Kapp(Not,[t]) -> - let d = decidability t in + | Kapp(Not,[t]) -> + let d = decidability t in tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]); + (generalize_tac + [mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]); (onClearedName i (fun i -> (loop (LocalAssum (make_annot i Sorts.Relevant,t) :: lit)))) ] - | Kapp(op,[t1;t2]) -> - (try - let thm = not_binop op in + | Kapp(op,[t1;t2]) -> + (try + let thm = not_binop op in tclTHENLIST [ - (generalize_tac - [mkApp (Lazy.force thm, [| t1;t2;mkVar i|])]); - (onClearedName i (fun _ -> loop lit)) + (generalize_tac + [mkApp (Lazy.force thm, [| t1;t2;mkVar i|])]); + (onClearedName i (fun _ -> loop lit)) ] - with Not_found -> loop lit) - | Kapp(Eq,[typ;t1;t2]) -> + with Not_found -> loop lit) + | Kapp(Eq,[typ;t1;t2]) -> if !old_style_flag then begin match destructurate_type env sigma typ with - | Kapp(Nat,_) -> + | Kapp(Nat,_) -> tclTHENLIST [ - (simplest_elim - (mkApp + (simplest_elim + (mkApp (Lazy.force coq_not_eq, [|t1;t2;mkVar i|]))); - (onClearedName i (fun _ -> loop lit)) + (onClearedName i (fun _ -> loop lit)) ] - | Kapp(Z,_) -> + | Kapp(Z,_) -> tclTHENLIST [ - (simplest_elim - (mkApp - (Lazy.force coq_not_Zeq, [|t1;t2;mkVar i|]))); - (onClearedName i (fun _ -> loop lit)) + (simplest_elim + (mkApp + (Lazy.force coq_not_Zeq, [|t1;t2;mkVar i|]))); + (onClearedName i (fun _ -> loop lit)) ] - | _ -> loop lit + | _ -> loop lit end else begin match destructurate_type env sigma typ with - | Kapp(Nat,_) -> + | Kapp(Nat,_) -> (tclTHEN (Tactics.convert_hyp ~check:false ~reorder:false (NamedDecl.set_type (mkApp (Lazy.force coq_neq, [| t1;t2|])) decl)) - (loop lit)) - | Kapp(Z,_) -> + (loop lit)) + | Kapp(Z,_) -> (tclTHEN (Tactics.convert_hyp ~check:false ~reorder:false (NamedDecl.set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|])) decl)) - (loop lit)) - | _ -> loop lit + (loop lit)) + | _ -> loop lit end - | _ -> loop lit + | _ -> loop lit end | _ -> loop lit - with - | Undecidable -> loop lit - | e when catchable_exception e -> loop lit - end + with + | Undecidable -> loop lit + | e when catchable_exception e -> loop lit + end in let hyps = Proofview.Goal.hyps gl in loop hyps @@ -1883,23 +1883,23 @@ let destructure_goal = match prop with | Kapp(Not,[t]) -> (tclTHEN - (tclTHEN (unfold sp_not) intro) - destructure_hyps) + (tclTHEN (unfold sp_not) intro) + destructure_hyps) | Kimp(a,b) -> (tclTHEN intro (loop b)) | Kapp(False,[]) -> destructure_hyps | _ -> - let goal_tac = - try - let dec = decidability t in - tclTHEN + let goal_tac = + try + let dec = decidability t in + tclTHEN (Proofview.Goal.enter begin fun gl -> - refine_app gl (mkApp (Lazy.force coq_dec_not_not, [| t; dec |])) - end) - intro - with Undecidable -> Tactics.elim_type (Lazy.force coq_False) - | e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e - in - tclTHEN goal_tac destructure_hyps + refine_app gl (mkApp (Lazy.force coq_dec_not_not, [| t; dec |])) + end) + intro + with Undecidable -> Tactics.elim_type (Lazy.force coq_False) + | e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e + in + tclTHEN goal_tac destructure_hyps in (loop concl) end diff --git a/plugins/omega/omega.ml b/plugins/omega/omega.ml index 05c31062fc..355e61deb9 100644 --- a/plugins/omega/omega.ml +++ b/plugins/omega/omega.ml @@ -130,14 +130,14 @@ let display_eq print_var (l,e) = let _ = List.fold_left (fun not_first f -> - print_string - (if f.c <? zero then "- " else if not_first then "+ " else ""); - let c = abs f.c in - if c =? one then - Printf.printf "%s " (print_var f.v) - else - Printf.printf "%s %s " (string_of_bigint c) (print_var f.v); - true) + print_string + (if f.c <? zero then "- " else if not_first then "+ " else ""); + let c = abs f.c in + if c =? one then + Printf.printf "%s " (print_var f.v) + else + Printf.printf "%s %s " (string_of_bigint c) (print_var f.v); + true) false l in if e >? zero then @@ -148,7 +148,7 @@ let display_eq print_var (l,e) = let rec trace_length l = let action_length accu = function | SPLIT_INEQ (_,(_,l1),(_,l2)) -> - accu + one + trace_length l1 + trace_length l2 + accu + one + trace_length l1 + trace_length l2 | _ -> accu + one in List.fold_left action_length zero l @@ -263,12 +263,12 @@ let rec sum p0 p1 = match (p0,p1) with | ([], l) -> l | (l, []) -> l | (((x1::l1) as l1'), ((x2::l2) as l2')) -> if x1.v = x2.v then - let c = x1.c + x2.c in - if c =? zero then sum l1 l2 else {v=x1.v;c=c} :: sum l1 l2 + let c = x1.c + x2.c in + if c =? zero then sum l1 l2 else {v=x1.v;c=c} :: sum l1 l2 else if x1.v > x2.v then - x1 :: sum l1 l2' + x1 :: sum l1 l2' else - x2 :: sum l1' l2 + x2 :: sum l1' l2 let sum_afine new_eq_id eq1 eq2 = { kind = eq1.kind; id = new_eq_id (); @@ -351,7 +351,7 @@ let banerjee_step (new_eq_id,new_var_id,print_var) original l1 l2 = original.body; id = new_eq_id (); kind = EQUA } in add_event (STATE {st_new_eq = new_eq; st_def = definition; - st_orig = original; st_coef = m; st_var = sigma}); + st_orig = original; st_coef = m; st_var = sigma}); let new_eq = List.hd (normalize new_eq) in let eliminated_var, def = chop_var var new_eq.body in let other_equations = @@ -395,8 +395,8 @@ let rec banerjee ((_,_,print_var) as new_ids) (sys_eq,sys_ineq) = add_event (CONSTANT_NOT_NUL(eq.id,eq.constant)); raise UNSOLVABLE end else - banerjee new_ids - (eliminate_one_equation new_ids (eq,other,sys_ineq)) + banerjee new_ids + (eliminate_one_equation new_ids (eq,other,sys_ineq)) type kind = INVERTED | NORMAL @@ -501,7 +501,7 @@ let product new_eq_id dark_shadow low high = (map_eq_afine (fun c -> c * a) eq2) in add_event(SUM(eq.id,(b,eq1),(a,eq2))); match normalize eq with - | [eq] -> + | [eq] -> let final_eq = if dark_shadow then let delta = (a - one) * (b - one) in @@ -549,43 +549,43 @@ let simplify ((new_eq_id,new_var_id,print_var) as new_ids) dark_shadow system = let rec depend relie_on accu = function | act :: l -> begin match act with - | DIVIDE_AND_APPROX (e,_,_,_) -> + | DIVIDE_AND_APPROX (e,_,_,_) -> if Int.List.mem e.id relie_on then depend relie_on (act::accu) l else depend relie_on accu l - | EXACT_DIVIDE (e,_) -> + | EXACT_DIVIDE (e,_) -> if Int.List.mem e.id relie_on then depend relie_on (act::accu) l else depend relie_on accu l - | WEAKEN (e,_) -> + | WEAKEN (e,_) -> if Int.List.mem e relie_on then depend relie_on (act::accu) l else depend relie_on accu l - | SUM (e,(_,e1),(_,e2)) -> + | SUM (e,(_,e1),(_,e2)) -> if Int.List.mem e relie_on then - depend (e1.id::e2.id::relie_on) (act::accu) l + depend (e1.id::e2.id::relie_on) (act::accu) l else - depend relie_on accu l - | STATE {st_new_eq=e;st_orig=o} -> + depend relie_on accu l + | STATE {st_new_eq=e;st_orig=o} -> if Int.List.mem e.id relie_on then depend (o.id::relie_on) (act::accu) l else depend relie_on accu l - | HYP e -> + | HYP e -> if Int.List.mem e.id relie_on then depend relie_on (act::accu) l else depend relie_on accu l - | FORGET_C _ -> depend relie_on accu l - | FORGET _ -> depend relie_on accu l - | FORGET_I _ -> depend relie_on accu l - | MERGE_EQ (e,e1,e2) -> + | FORGET_C _ -> depend relie_on accu l + | FORGET _ -> depend relie_on accu l + | FORGET_I _ -> depend relie_on accu l + | MERGE_EQ (e,e1,e2) -> if Int.List.mem e relie_on then - depend (e1.id::e2::relie_on) (act::accu) l + depend (e1.id::e2::relie_on) (act::accu) l else - depend relie_on accu l - | NOT_EXACT_DIVIDE (e,_) -> depend (e.id::relie_on) (act::accu) l - | CONTRADICTION (e1,e2) -> - depend (e1.id::e2.id::relie_on) (act::accu) l - | CONSTANT_NOT_NUL (e,_) -> depend (e::relie_on) (act::accu) l - | CONSTANT_NEG (e,_) -> depend (e::relie_on) (act::accu) l - | CONSTANT_NUL e -> depend (e::relie_on) (act::accu) l - | NEGATE_CONTRADICT (e1,e2,_) -> + depend relie_on accu l + | NOT_EXACT_DIVIDE (e,_) -> depend (e.id::relie_on) (act::accu) l + | CONTRADICTION (e1,e2) -> depend (e1.id::e2.id::relie_on) (act::accu) l - | SPLIT_INEQ _ -> failwith "depend" + | CONSTANT_NOT_NUL (e,_) -> depend (e::relie_on) (act::accu) l + | CONSTANT_NEG (e,_) -> depend (e::relie_on) (act::accu) l + | CONSTANT_NUL e -> depend (e::relie_on) (act::accu) l + | NEGATE_CONTRADICT (e1,e2,_) -> + depend (e1.id::e2.id::relie_on) (act::accu) l + | SPLIT_INEQ _ -> failwith "depend" end | [] -> relie_on, accu @@ -602,9 +602,9 @@ let negation (eqs,ineqs) = assert (e.kind = EQUA); let {body=ne;constant=c},kind = normal e in try - let (kind',e') = Hashtbl.find table (ne,c) in - add_event (NEGATE_CONTRADICT (e,e',kind=kind')); - raise UNSOLVABLE + let (kind',e') = Hashtbl.find table (ne,c) in + add_event (NEGATE_CONTRADICT (e,e',kind=kind')); + raise UNSOLVABLE with Not_found -> ()) eqs exception FULL_SOLUTION of action list * int list @@ -631,20 +631,20 @@ let simplify_strong ((new_eq_id,new_var_id,print_var) as new_ids) system = in let rec explode_diseq = function | (de::diseq,ineqs,expl_map) -> - let id1 = new_eq_id () - and id2 = new_eq_id () in - let e1 = + let id1 = new_eq_id () + and id2 = new_eq_id () in + let e1 = {id = id1; kind=INEQ; body = de.body; constant = de.constant -one} in - let e2 = - {id = id2; kind=INEQ; body = map_eq_linear neg de.body; + let e2 = + {id = id2; kind=INEQ; body = map_eq_linear neg de.body; constant = neg de.constant - one} in - let new_sys = + let new_sys = List.map (fun (what,sys) -> ((de.id,id1,true)::what, e1::sys)) - ineqs @ + ineqs @ List.map (fun (what,sys) -> ((de.id,id2,false)::what,e2::sys)) - ineqs - in - explode_diseq (diseq,new_sys,(de.id,(de,id1,id2))::expl_map) + ineqs + in + explode_diseq (diseq,new_sys,(de.id,(de,id1,id2))::expl_map) | ([],ineqs,expl_map) -> ineqs,expl_map in try @@ -673,19 +673,19 @@ let simplify_strong ((new_eq_id,new_var_id,print_var) as new_ids) system = let tbl = Hashtbl.create 7 in let augment x = try incr (Hashtbl.find tbl x) - with Not_found -> Hashtbl.add tbl x (ref 1) in + with Not_found -> Hashtbl.add tbl x (ref 1) in let eq = ref (-1) and c = ref 0 in List.iter (function - | ([],r_on,_,path) -> raise (FULL_SOLUTION (path,r_on)) + | ([],r_on,_,path) -> raise (FULL_SOLUTION (path,r_on)) | (l,_,_,_) -> List.iter augment l) sys; Hashtbl.iter (fun x v -> if !v > !c then begin eq := x; c := !v end) tbl; !eq in let rec solve systems = try - let id = max_count systems in + let id = max_count systems in let rec sign = function - | ((id',_,b)::l) -> if id=id' then b else sign l + | ((id',_,b)::l) -> if id=id' then b else sign l | [] -> failwith "solve" in let s1,s2 = List.partition (fun (_,_,decomp,_) -> sign decomp) systems in @@ -695,9 +695,9 @@ let simplify_strong ((new_eq_id,new_var_id,print_var) as new_ids) system = let s1' = List.map remove_int s1 in let s2' = List.map remove_int s2 in let (r1,relie1) = solve s1' - and (r2,relie2) = solve s2' in - let (eq,id1,id2) = Int.List.assoc id explode_map in - [SPLIT_INEQ(eq,(id1,r1),(id2, r2))], + and (r2,relie2) = solve s2' in + let (eq,id1,id2) = Int.List.assoc id explode_map in + [SPLIT_INEQ(eq,(id1,r1),(id2, r2))], eq.id :: Util.List.union Int.equal relie1 relie2 with FULL_SOLUTION (x0,x1) -> (x0,x1) in diff --git a/plugins/rtauto/proof_search.ml b/plugins/rtauto/proof_search.ml index 13da8220f4..4cc32cfb26 100644 --- a/plugins/rtauto/proof_search.ml +++ b/plugins/rtauto/proof_search.ml @@ -102,7 +102,7 @@ type sequent = let add_one_arrow i f1 f2 m= try Fmap.add f1 ((i,f2)::(Fmap.find f1 m)) m with Not_found -> - Fmap.add f1 [i,f2] m + Fmap.add f1 [i,f2] m type proof = Ax of int @@ -174,7 +174,7 @@ let project = function let pop n prf = let nprf= match prf.dep_it with - Pop (i,p) -> Pop (i+n,p) + Pop (i,p) -> Pop (i+n,p) | p -> Pop(n,p) in {prf with dep_it = nprf} @@ -182,71 +182,71 @@ let rec fill stack proof = match stack with [] -> Complete proof.dep_it | slice::super -> - if - !pruning && - List.is_empty slice.proofs_done && - not (slice.changes_goal && proof.dep_goal) && - not (Int.Set.exists - (fun i -> Int.Set.mem i proof.dep_hyps) - slice.creates_hyps) - then - begin - s_info.pruned_steps<-s_info.pruned_steps+1; - s_info.pruned_branches<- s_info.pruned_branches + - List.length slice.proofs_todo; - let created_here=Int.Set.cardinal slice.creates_hyps in - s_info.pruned_hyps<-s_info.pruned_hyps+ - List.fold_left - (fun sum dseq -> sum + Int.Set.cardinal dseq.dep_hyps) - created_here slice.proofs_todo; - fill super (pop (Int.Set.cardinal slice.creates_hyps) proof) - end - else - let dep_hyps= - Int.Set.union slice.needs_hyps - (Int.Set.diff proof.dep_hyps slice.creates_hyps) in - let dep_goal= - slice.needs_goal || - ((not slice.changes_goal) && proof.dep_goal) in - let proofs_done= - proof.dep_it::slice.proofs_done in - match slice.proofs_todo with - [] -> - fill super {dep_it = - add_step slice.step (List.rev proofs_done); - dep_goal = dep_goal; - dep_hyps = dep_hyps} - | current::next -> - let nslice= - {proofs_done=proofs_done; - proofs_todo=next; - step=slice.step; - needs_goal=dep_goal; - needs_hyps=dep_hyps; - changes_goal=current.dep_goal; - creates_hyps=current.dep_hyps} in - Incomplete (current.dep_it,nslice::super) + if + !pruning && + List.is_empty slice.proofs_done && + not (slice.changes_goal && proof.dep_goal) && + not (Int.Set.exists + (fun i -> Int.Set.mem i proof.dep_hyps) + slice.creates_hyps) + then + begin + s_info.pruned_steps<-s_info.pruned_steps+1; + s_info.pruned_branches<- s_info.pruned_branches + + List.length slice.proofs_todo; + let created_here=Int.Set.cardinal slice.creates_hyps in + s_info.pruned_hyps<-s_info.pruned_hyps+ + List.fold_left + (fun sum dseq -> sum + Int.Set.cardinal dseq.dep_hyps) + created_here slice.proofs_todo; + fill super (pop (Int.Set.cardinal slice.creates_hyps) proof) + end + else + let dep_hyps= + Int.Set.union slice.needs_hyps + (Int.Set.diff proof.dep_hyps slice.creates_hyps) in + let dep_goal= + slice.needs_goal || + ((not slice.changes_goal) && proof.dep_goal) in + let proofs_done= + proof.dep_it::slice.proofs_done in + match slice.proofs_todo with + [] -> + fill super {dep_it = + add_step slice.step (List.rev proofs_done); + dep_goal = dep_goal; + dep_hyps = dep_hyps} + | current::next -> + let nslice= + {proofs_done=proofs_done; + proofs_todo=next; + step=slice.step; + needs_goal=dep_goal; + needs_hyps=dep_hyps; + changes_goal=current.dep_goal; + creates_hyps=current.dep_hyps} in + Incomplete (current.dep_it,nslice::super) let append stack (step,subgoals) = s_info.created_steps<-s_info.created_steps+1; match subgoals with [] -> - s_info.branch_successes<-s_info.branch_successes+1; - fill stack {dep_it=add_step step.dep_it []; - dep_goal=step.dep_goal; - dep_hyps=step.dep_hyps} + s_info.branch_successes<-s_info.branch_successes+1; + fill stack {dep_it=add_step step.dep_it []; + dep_goal=step.dep_goal; + dep_hyps=step.dep_hyps} | hd :: next -> - s_info.created_branches<- - s_info.created_branches+List.length next; - let slice= - {proofs_done=[]; - proofs_todo=next; - step=step.dep_it; - needs_goal=step.dep_goal; - needs_hyps=step.dep_hyps; - changes_goal=hd.dep_goal; - creates_hyps=hd.dep_hyps} in - Incomplete(hd.dep_it,slice::stack) + s_info.created_branches<- + s_info.created_branches+List.length next; + let slice= + {proofs_done=[]; + proofs_todo=next; + step=step.dep_it; + needs_goal=step.dep_goal; + needs_hyps=step.dep_hyps; + changes_goal=hd.dep_goal; + creates_hyps=hd.dep_hyps} in + Incomplete(hd.dep_it,slice::stack) let embed seq= {dep_it=seq; @@ -266,59 +266,59 @@ let add_hyp seqwd f= let cnx,right= try let l=Fmap.find f seq.right in - List.fold_right (fun (i,f0) l0 -> (num,i,f,f0)::l0) l seq.cnx, - Fmap.remove f seq.right + List.fold_right (fun (i,f0) l0 -> (num,i,f,f0)::l0) l seq.cnx, + Fmap.remove f seq.right with Not_found -> seq.cnx,seq.right in let nseq= match f with - Bot -> - {seq with - left=left; - right=right; - size=num; - abs=Some num; - cnx=cnx} + Bot -> + {seq with + left=left; + right=right; + size=num; + abs=Some num; + cnx=cnx} | Atom _ -> - {seq with - size=num; - left=left; - right=right; - cnx=cnx} + {seq with + size=num; + left=left; + right=right; + cnx=cnx} | Conjunct (_,_) | Disjunct (_,_) -> - {seq with - rev_hyps=Int.Map.add num f seq.rev_hyps; - size=num; - left=left; - right=right; - cnx=cnx} + {seq with + rev_hyps=Int.Map.add num f seq.rev_hyps; + size=num; + left=left; + right=right; + cnx=cnx} | Arrow (f1,f2) -> - let ncnx,nright= - try - let i = Fmap.find f1 seq.left in - (i,num,f1,f2)::cnx,right - with Not_found -> - cnx,(add_one_arrow num f1 f2 right) in - match f1 with - Conjunct (_,_) | Disjunct (_,_) -> - {seq with - rev_hyps=Int.Map.add num f seq.rev_hyps; - size=num; - left=left; - right=nright; - cnx=ncnx} - | Arrow(_,_) -> - {seq with - norev_hyps=Int.Map.add num f seq.norev_hyps; - size=num; - left=left; - right=nright; - cnx=ncnx} - | _ -> - {seq with - size=num; - left=left; - right=nright; - cnx=ncnx} in + let ncnx,nright= + try + let i = Fmap.find f1 seq.left in + (i,num,f1,f2)::cnx,right + with Not_found -> + cnx,(add_one_arrow num f1 f2 right) in + match f1 with + Conjunct (_,_) | Disjunct (_,_) -> + {seq with + rev_hyps=Int.Map.add num f seq.rev_hyps; + size=num; + left=left; + right=nright; + cnx=ncnx} + | Arrow(_,_) -> + {seq with + norev_hyps=Int.Map.add num f seq.norev_hyps; + size=num; + left=left; + right=nright; + cnx=ncnx} + | _ -> + {seq with + size=num; + left=left; + right=nright; + cnx=ncnx} in {seqwd with dep_it=nseq; dep_hyps=Int.Set.add num seqwd.dep_hyps} @@ -336,33 +336,33 @@ let choose m= let search_or seq= match seq.gl with Disjunct (f1,f2) -> - [{dep_it = SI_Or_l; - dep_goal = true; - dep_hyps = Int.Set.empty}, - [change_goal (embed seq) f1]; - {dep_it = SI_Or_r; - dep_goal = true; - dep_hyps = Int.Set.empty}, - [change_goal (embed seq) f2]] + [{dep_it = SI_Or_l; + dep_goal = true; + dep_hyps = Int.Set.empty}, + [change_goal (embed seq) f1]; + {dep_it = SI_Or_r; + dep_goal = true; + dep_hyps = Int.Set.empty}, + [change_goal (embed seq) f2]] | _ -> [] let search_norev seq= let goals=ref (search_or seq) in let add_one i f= match f with - Arrow (Arrow (f1,f2),f3) -> - let nseq = - {seq with norev_hyps=Int.Map.remove i seq.norev_hyps} in - goals:= - ({dep_it=SD_Arrow(i); - dep_goal=false; - dep_hyps=Int.Set.singleton i}, - [add_hyp - (add_hyp - (change_goal (embed nseq) f2) - (Arrow(f2,f3))) - f1; - add_hyp (embed nseq) f3]):: !goals + Arrow (Arrow (f1,f2),f3) -> + let nseq = + {seq with norev_hyps=Int.Map.remove i seq.norev_hyps} in + goals:= + ({dep_it=SD_Arrow(i); + dep_goal=false; + dep_hyps=Int.Set.singleton i}, + [add_hyp + (add_hyp + (change_goal (embed nseq) f2) + (Arrow(f2,f3))) + f1; + add_hyp (embed nseq) f3]):: !goals | _ -> anomaly ~label:"search_no_rev" (Pp.str "can't happen.") in Int.Map.iter add_one seq.norev_hyps; List.rev !goals @@ -376,76 +376,76 @@ let search_in_rev_hyps seq= dep_hyps=Int.Set.singleton i} in let nseq={seq with rev_hyps=Int.Map.remove i seq.rev_hyps} in match f with - Conjunct (f1,f2) -> - [make_step (SE_And(i)), - [add_hyp (add_hyp (embed nseq) f1) f2]] - | Disjunct (f1,f2) -> - [make_step (SE_Or(i)), - [add_hyp (embed nseq) f1;add_hyp (embed nseq) f2]] - | Arrow (Conjunct (f1,f2),f0) -> - [make_step (SD_And(i)), - [add_hyp (embed nseq) (Arrow (f1,Arrow (f2,f0)))]] - | Arrow (Disjunct (f1,f2),f0) -> - [make_step (SD_Or(i)), - [add_hyp (add_hyp (embed nseq) (Arrow(f1,f0))) (Arrow (f2,f0))]] - | _ -> anomaly ~label:"search_in_rev_hyps" (Pp.str "can't happen.") + Conjunct (f1,f2) -> + [make_step (SE_And(i)), + [add_hyp (add_hyp (embed nseq) f1) f2]] + | Disjunct (f1,f2) -> + [make_step (SE_Or(i)), + [add_hyp (embed nseq) f1;add_hyp (embed nseq) f2]] + | Arrow (Conjunct (f1,f2),f0) -> + [make_step (SD_And(i)), + [add_hyp (embed nseq) (Arrow (f1,Arrow (f2,f0)))]] + | Arrow (Disjunct (f1,f2),f0) -> + [make_step (SD_Or(i)), + [add_hyp (add_hyp (embed nseq) (Arrow(f1,f0))) (Arrow (f2,f0))]] + | _ -> anomaly ~label:"search_in_rev_hyps" (Pp.str "can't happen.") with Not_found -> search_norev seq let search_rev seq= match seq.cnx with (i,j,f1,f2)::next -> - let nseq= - match f1 with - Conjunct (_,_) | Disjunct (_,_) -> - {seq with cnx=next; - rev_hyps=Int.Map.remove j seq.rev_hyps} - | Arrow (_,_) -> - {seq with cnx=next; - norev_hyps=Int.Map.remove j seq.norev_hyps} - | _ -> - {seq with cnx=next} in - [{dep_it=SE_Arrow(i,j); - dep_goal=false; - dep_hyps=Int.Set.add i (Int.Set.singleton j)}, - [add_hyp (embed nseq) f2]] + let nseq= + match f1 with + Conjunct (_,_) | Disjunct (_,_) -> + {seq with cnx=next; + rev_hyps=Int.Map.remove j seq.rev_hyps} + | Arrow (_,_) -> + {seq with cnx=next; + norev_hyps=Int.Map.remove j seq.norev_hyps} + | _ -> + {seq with cnx=next} in + [{dep_it=SE_Arrow(i,j); + dep_goal=false; + dep_hyps=Int.Set.add i (Int.Set.singleton j)}, + [add_hyp (embed nseq) f2]] | [] -> - match seq.gl with - Arrow (f1,f2) -> - [{dep_it=SI_Arrow; - dep_goal=true; - dep_hyps=Int.Set.empty}, - [add_hyp (change_goal (embed seq) f2) f1]] - | Conjunct (f1,f2) -> - [{dep_it=SI_And; - dep_goal=true; - dep_hyps=Int.Set.empty},[change_goal (embed seq) f1; - change_goal (embed seq) f2]] - | _ -> search_in_rev_hyps seq + match seq.gl with + Arrow (f1,f2) -> + [{dep_it=SI_Arrow; + dep_goal=true; + dep_hyps=Int.Set.empty}, + [add_hyp (change_goal (embed seq) f2) f1]] + | Conjunct (f1,f2) -> + [{dep_it=SI_And; + dep_goal=true; + dep_hyps=Int.Set.empty},[change_goal (embed seq) f1; + change_goal (embed seq) f2]] + | _ -> search_in_rev_hyps seq let search_all seq= match seq.abs with Some i -> - [{dep_it=SE_False (i); - dep_goal=false; - dep_hyps=Int.Set.singleton i},[]] + [{dep_it=SE_False (i); + dep_goal=false; + dep_hyps=Int.Set.singleton i},[]] | None -> - try - let ax = Fmap.find seq.gl seq.left in - [{dep_it=SAx (ax); - dep_goal=true; - dep_hyps=Int.Set.singleton ax},[]] - with Not_found -> search_rev seq + try + let ax = Fmap.find seq.gl seq.left in + [{dep_it=SAx (ax); + dep_goal=true; + dep_hyps=Int.Set.singleton ax},[]] + with Not_found -> search_rev seq let bare_sequent = embed - {rev_hyps=Int.Map.empty; - norev_hyps=Int.Map.empty; - size=0; - left=Fmap.empty; - right=Fmap.empty; - cnx=[]; - abs=None; - gl=Bot} + {rev_hyps=Int.Map.empty; + norev_hyps=Int.Map.empty; + size=0; + left=Fmap.empty; + right=Fmap.empty; + cnx=[]; + abs=None; + gl=Bot} let init_state hyps gl= let init = change_goal bare_sequent gl in @@ -461,11 +461,11 @@ let branching = function Control.check_for_interrupt (); let successors = search_all seq in let _ = - match successors with - [] -> s_info.branch_failures<-s_info.branch_failures+1 - | _::next -> - s_info.nd_branching<-s_info.nd_branching+List.length next in - List.map (append stack) successors + match successors with + [] -> s_info.branch_failures<-s_info.branch_failures+1 + | _::next -> + s_info.nd_branching<-s_info.nd_branching+List.length next in + List.map (append stack) successors | Complete prf -> anomaly (Pp.str "already succeeded.") open Pp @@ -492,7 +492,7 @@ let pr_form f = pp_form f let pp_intmap map = let pp=ref (str "") in Int.Map.iter (fun i obj -> pp:= (!pp ++ - pp_form obj ++ cut ())) map; + pp_form obj ++ cut ())) map; str "{ " ++ v 0 (!pp) ++ str " }" let pp_list pp_obj l= @@ -503,20 +503,20 @@ let pp=ref (str "") in let pp_mapint map = let pp=ref (str "") in Fmap.iter (fun obj l -> pp:= (!pp ++ - pp_form obj ++ str " => " ++ - pp_list (fun (i,f) -> pp_form f) l ++ - cut ()) ) map; + pp_form obj ++ str " => " ++ + pp_list (fun (i,f) -> pp_form f) l ++ + cut ()) ) map; str "{ " ++ hv 0 (!pp ++ str " }") let pp_connect (i,j,f1,f2) = pp_form f1 ++ str " => " ++ pp_form f2 let pp_gl gl= cut () ++ str "{ " ++ hv 0 ( - begin - match gl.abs with - None -> str "" - | Some i -> str "ABSURD" ++ cut () - end ++ + begin + match gl.abs with + None -> str "" + | Some i -> str "ABSURD" ++ cut () + end ++ str "rev =" ++ pp_intmap gl.rev_hyps ++ cut () ++ str "norev =" ++ pp_intmap gl.norev_hyps ++ cut () ++ str "arrows=" ++ pp_mapint gl.right ++ cut () ++ @@ -531,31 +531,31 @@ let pp = let pp_info () = let count_info = if !pruning then - str "Proof steps : " ++ - int s_info.created_steps ++ str " created / " ++ - int s_info.pruned_steps ++ str " pruned" ++ fnl () ++ - str "Proof branches : " ++ - int s_info.created_branches ++ str " created / " ++ - int s_info.pruned_branches ++ str " pruned" ++ fnl () ++ - str "Hypotheses : " ++ - int s_info.created_hyps ++ str " created / " ++ - int s_info.pruned_hyps ++ str " pruned" ++ fnl () + str "Proof steps : " ++ + int s_info.created_steps ++ str " created / " ++ + int s_info.pruned_steps ++ str " pruned" ++ fnl () ++ + str "Proof branches : " ++ + int s_info.created_branches ++ str " created / " ++ + int s_info.pruned_branches ++ str " pruned" ++ fnl () ++ + str "Hypotheses : " ++ + int s_info.created_hyps ++ str " created / " ++ + int s_info.pruned_hyps ++ str " pruned" ++ fnl () else - str "Pruning is off" ++ fnl () ++ - str "Proof steps : " ++ - int s_info.created_steps ++ str " created" ++ fnl () ++ - str "Proof branches : " ++ - int s_info.created_branches ++ str " created" ++ fnl () ++ - str "Hypotheses : " ++ - int s_info.created_hyps ++ str " created" ++ fnl () in + str "Pruning is off" ++ fnl () ++ + str "Proof steps : " ++ + int s_info.created_steps ++ str " created" ++ fnl () ++ + str "Proof branches : " ++ + int s_info.created_branches ++ str " created" ++ fnl () ++ + str "Hypotheses : " ++ + int s_info.created_hyps ++ str " created" ++ fnl () in Feedback.msg_info ( str "Proof-search statistics :" ++ fnl () ++ - count_info ++ - str "Branch ends: " ++ - int s_info.branch_successes ++ str " successes / " ++ - int s_info.branch_failures ++ str " failures" ++ fnl () ++ - str "Non-deterministic choices : " ++ - int s_info.nd_branching ++ str " branches") + count_info ++ + str "Branch ends: " ++ + int s_info.branch_successes ++ str " successes / " ++ + int s_info.branch_failures ++ str " failures" ++ fnl () ++ + str "Non-deterministic choices : " ++ + int s_info.nd_branching ++ str " branches") diff --git a/plugins/rtauto/refl_tauto.ml b/plugins/rtauto/refl_tauto.ml index df27c9c9d7..0c155c9d0a 100644 --- a/plugins/rtauto/refl_tauto.ml +++ b/plugins/rtauto/refl_tauto.ml @@ -156,9 +156,9 @@ let rec decal k = function [] -> k | (start,delta)::rest -> if k>start then - k - delta + k - delta else - decal k rest + decal k rest let add_pop size d pops= match pops with @@ -168,57 +168,57 @@ let add_pop size d pops= let rec build_proof pops size = function Ax i -> - mkApp (force step_count l_Ax, - [|build_pos (decal i pops)|]) + mkApp (force step_count l_Ax, + [|build_pos (decal i pops)|]) | I_Arrow p -> - mkApp (force step_count l_I_Arrow, - [|build_proof pops (size + 1) p|]) + mkApp (force step_count l_I_Arrow, + [|build_proof pops (size + 1) p|]) | E_Arrow(i,j,p) -> - mkApp (force step_count l_E_Arrow, - [|build_pos (decal i pops); - build_pos (decal j pops); - build_proof pops (size + 1) p|]) + mkApp (force step_count l_E_Arrow, + [|build_pos (decal i pops); + build_pos (decal j pops); + build_proof pops (size + 1) p|]) | D_Arrow(i,p1,p2) -> - mkApp (force step_count l_D_Arrow, - [|build_pos (decal i pops); - build_proof pops (size + 2) p1; - build_proof pops (size + 1) p2|]) + mkApp (force step_count l_D_Arrow, + [|build_pos (decal i pops); + build_proof pops (size + 2) p1; + build_proof pops (size + 1) p2|]) | E_False i -> - mkApp (force step_count l_E_False, - [|build_pos (decal i pops)|]) + mkApp (force step_count l_E_False, + [|build_pos (decal i pops)|]) | I_And(p1,p2) -> - mkApp (force step_count l_I_And, - [|build_proof pops size p1; - build_proof pops size p2|]) + mkApp (force step_count l_I_And, + [|build_proof pops size p1; + build_proof pops size p2|]) | E_And(i,p) -> - mkApp (force step_count l_E_And, - [|build_pos (decal i pops); - build_proof pops (size + 2) p|]) + mkApp (force step_count l_E_And, + [|build_pos (decal i pops); + build_proof pops (size + 2) p|]) | D_And(i,p) -> - mkApp (force step_count l_D_And, - [|build_pos (decal i pops); - build_proof pops (size + 1) p|]) + mkApp (force step_count l_D_And, + [|build_pos (decal i pops); + build_proof pops (size + 1) p|]) | I_Or_l(p) -> - mkApp (force step_count l_I_Or_l, - [|build_proof pops size p|]) + mkApp (force step_count l_I_Or_l, + [|build_proof pops size p|]) | I_Or_r(p) -> - mkApp (force step_count l_I_Or_r, - [|build_proof pops size p|]) + mkApp (force step_count l_I_Or_r, + [|build_proof pops size p|]) | E_Or(i,p1,p2) -> - mkApp (force step_count l_E_Or, - [|build_pos (decal i pops); - build_proof pops (size + 1) p1; - build_proof pops (size + 1) p2|]) + mkApp (force step_count l_E_Or, + [|build_pos (decal i pops); + build_proof pops (size + 1) p1; + build_proof pops (size + 1) p2|]) | D_Or(i,p) -> - mkApp (force step_count l_D_Or, - [|build_pos (decal i pops); - build_proof pops (size + 2) p|]) + mkApp (force step_count l_D_Or, + [|build_pos (decal i pops); + build_proof pops (size + 2) p|]) | Pop(d,p) -> - build_proof (add_pop size d pops) size p + build_proof (add_pop size d pops) size p let build_env gamma= List.fold_right (fun (p,_) e -> - mkApp(force node_count l_push,[|mkProp;p;e|])) + mkApp(force node_count l_push,[|mkProp;p;e|])) gamma.env (mkApp (force node_count l_empty,[|mkProp|])) open Goptions diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml index e3e787df2c..f1dc63dd9e 100644 --- a/plugins/setoid_ring/newring.ml +++ b/plugins/setoid_ring/newring.ml @@ -41,12 +41,12 @@ type protect_flag = Eval|Prot|Rec type protection = Evd.evar_map -> EConstr.t -> GlobRef.t -> (Int.t -> protect_flag) option -let global_head_of_constr sigma c = +let global_head_of_constr sigma c = let f, args = decompose_app sigma c in try fst (Termops.global_of_constr sigma f) with Not_found -> CErrors.anomaly (str "global_head_of_constr.") -let global_of_constr_nofail c = +let global_of_constr_nofail c = try global_of_constr c with Not_found -> GlobRef.VarRef (Id.of_string "dummy") @@ -163,7 +163,7 @@ let ltac_call tac (args:glob_tactic_arg list) = TacArg(CAst.make @@ TacCall (CAst.make (ArgArg(Loc.tag @@ Lazy.force tac),args))) let dummy_goal env sigma = - let (gl,_,sigma) = + let (gl,_,sigma) = Goal.V82.mk_goal sigma (named_context_val env) EConstr.mkProp in {Evd.it = gl; Evd.sigma = sigma} @@ -428,9 +428,9 @@ let op_smorph r add mul req m1 m2 = let ring_equality env evd (r,add,mul,opp,req) = match EConstr.kind !evd req with | App (f, [| _ |]) when eq_constr_nounivs !evd f (Lazy.force coq_eq) -> - let setoid = plapp evd coq_eq_setoid [|r|] in - let op_morph = - match opp with + let setoid = plapp evd coq_eq_setoid [|r|] in + let op_morph = + match opp with Some opp -> plapp evd coq_eq_morph [|r;add;mul;opp|] | None -> plapp evd coq_eq_smorph [|r;add;mul|] in let sigma = !evd in @@ -439,41 +439,41 @@ let ring_equality env evd (r,add,mul,opp,req) = evd := sigma; (setoid,op_morph) | _ -> - let setoid = setoid_of_relation (Global.env ()) evd r req in - let signature = [Some (r,Some req);Some (r,Some req)],Some(r,Some req) in - let add_m, add_m_lem = - try Rewrite.default_morphism signature add + let setoid = setoid_of_relation (Global.env ()) evd r req in + let signature = [Some (r,Some req);Some (r,Some req)],Some(r,Some req) in + let add_m, add_m_lem = + try Rewrite.default_morphism signature add with Not_found -> error "ring addition should be declared as a morphism" in - let mul_m, mul_m_lem = + let mul_m, mul_m_lem = try Rewrite.default_morphism signature mul with Not_found -> error "ring multiplication should be declared as a morphism" in let op_morph = match opp with | Some opp -> - (let opp_m,opp_m_lem = - try Rewrite.default_morphism ([Some(r,Some req)],Some(r,Some req)) opp - with Not_found -> + (let opp_m,opp_m_lem = + try Rewrite.default_morphism ([Some(r,Some req)],Some(r,Some req)) opp + with Not_found -> error "ring opposite should be declared as a morphism" in - let op_morph = - op_morph r add mul opp req add_m_lem mul_m_lem opp_m_lem in - Flags.if_verbose - Feedback.msg_info + let op_morph = + op_morph r add mul opp req add_m_lem mul_m_lem opp_m_lem in + Flags.if_verbose + Feedback.msg_info (str"Using setoid \""++ pr_econstr_env env !evd req++str"\""++spc()++ str"and morphisms \""++pr_econstr_env env !evd add_m_lem ++ str"\","++spc()++ str"\""++pr_econstr_env env !evd mul_m_lem++ str"\""++spc()++str"and \""++pr_econstr_env env !evd opp_m_lem++ - str"\""); - op_morph) + str"\""); + op_morph) | None -> - (Flags.if_verbose - Feedback.msg_info + (Flags.if_verbose + Feedback.msg_info (str"Using setoid \""++pr_econstr_env env !evd req ++str"\"" ++ spc() ++ str"and morphisms \""++pr_econstr_env env !evd add_m_lem ++ - str"\""++spc()++str"and \""++ + str"\""++spc()++str"and \""++ pr_econstr_env env !evd mul_m_lem++str"\""); - op_smorph r add mul req add_m_lem mul_m_lem) in + op_smorph r add mul req add_m_lem mul_m_lem) in (setoid,op_morph) let build_setoid_params env evd r add mul opp req eqth = @@ -519,11 +519,11 @@ let make_hyp env evd c = let make_hyp_list env evdref lH = let evd, carrier = Evarutil.new_global !evdref (Lazy.force coq_hypo) in evdref := evd; - let l = + let l = List.fold_right (fun c l -> plapp evdref coq_cons [|carrier; (make_hyp env evdref c); l|]) lH (plapp evdref coq_nil [|carrier|]) - in + in let sigma, l' = Typing.solve_evars env !evdref l in evdref := sigma; let l' = EConstr.Unsafe.to_constr l' in @@ -609,7 +609,7 @@ let add_theory0 name (sigma, rth) eqth morphth cst_tac (pre,post) power sign div ring_morph = params.(2); ring_th = params.(0); ring_cst_tac = cst_tac; - ring_pow_tac = pow_tac; + ring_pow_tac = pow_tac; ring_lemma1 = lemma1; ring_lemma2 = lemma2; ring_pre_tac = pretac; @@ -867,13 +867,13 @@ let field_equality evd r inv req = let c = EConstr.of_constr c in mkApp(c,[|r;r;inv|]) | _ -> - let _setoid = setoid_of_relation (Global.env ()) evd r req in - let signature = [Some (r,Some req)],Some(r,Some req) in - let inv_m, inv_m_lem = - try Rewrite.default_morphism signature inv + let _setoid = setoid_of_relation (Global.env ()) evd r req in + let signature = [Some (r,Some req)],Some(r,Some req) in + let inv_m, inv_m_lem = + try Rewrite.default_morphism signature inv with Not_found -> error "field inverse should be declared as a morphism" in - inv_m_lem + inv_m_lem let add_field_theory0 name fth eqth morphth cst_tac inj (pre,post) power sign odiv = let open Constr in @@ -904,13 +904,13 @@ let add_field_theory0 name fth eqth morphth cst_tac inj (pre,post) power sign od | None -> params.(7) in let lemma1 = decl_constant (Id.to_string name^"_field_lemma1") ctx lemma1 in - let lemma2 = decl_constant (Id.to_string name^"_field_lemma2") + let lemma2 = decl_constant (Id.to_string name^"_field_lemma2") ctx lemma2 in - let lemma3 = decl_constant (Id.to_string name^"_field_lemma3") + let lemma3 = decl_constant (Id.to_string name^"_field_lemma3") ctx lemma3 in - let lemma4 = decl_constant (Id.to_string name^"_field_lemma4") + let lemma4 = decl_constant (Id.to_string name^"_field_lemma4") ctx lemma4 in - let cond_lemma = decl_constant (Id.to_string name^"_lemma5") + let cond_lemma = decl_constant (Id.to_string name^"_lemma5") ctx cond_lemma in let cst_tac = interp_cst_tac env sigma morphth kind (zero,one,add,mul,opp) cst_tac in diff --git a/plugins/ssr/ssrcommon.ml b/plugins/ssr/ssrcommon.ml index dbb60e6712..de3c660938 100644 --- a/plugins/ssr/ssrcommon.ml +++ b/plugins/ssr/ssrcommon.ml @@ -270,7 +270,7 @@ let of_ftactic ftac gl = in (sigma, ans) -let interp_wit wit ist gl x = +let interp_wit wit ist gl x = let globarg = in_gen (glbwit wit) x in let arg = Tacinterp.interp_genarg ist globarg in let (sigma, arg) = of_ftactic arg gl in @@ -350,7 +350,7 @@ let same_prefix s t n = let rec loop i = i = n || s.[i] = t.[i] && loop (i + 1) in loop 0 let skip_digits s = - let n = String.length s in + let n = String.length s in let rec loop i = if i < n && is_digit s.[i] then loop (i + 1) else i in loop let mk_tagged_id t i = Id.of_string (Printf.sprintf "%s%d_" t i) @@ -368,7 +368,7 @@ let wildcard_tag = "_the_" let wildcard_post = "_wildcard_" let mk_wildcard_id i = Id.of_string (Printf.sprintf "%s%s%s" wildcard_tag (CString.ordinal i) wildcard_post) -let has_wildcard_tag s = +let has_wildcard_tag s = let n = String.length s in let m = String.length wildcard_tag in let m' = String.length wildcard_post in n < m + m' + 2 && same_prefix s wildcard_tag m && @@ -440,7 +440,7 @@ let inc_safe n = if n = 0 then n else n + 1 let rec safe_depth s c = match EConstr.kind s c with | LetIn ({binder_name=Name x}, _, _, c') when is_discharged_id x -> safe_depth s c' + 1 | LetIn (_, _, _, c') | Prod (_, _, c') -> inc_safe (safe_depth s c') -| _ -> 0 +| _ -> 0 let red_safe (r : Reductionops.reduction_function) e s c0 = let rec red_to e c n = match EConstr.kind s c with @@ -518,7 +518,7 @@ let resolve_typeclasses ~where ~fail env sigma = sigma -let nf_evar sigma t = +let nf_evar sigma t = EConstr.Unsafe.to_constr (Evarutil.nf_evar sigma (EConstr.of_constr t)) let pf_abs_evars2 gl rigid (sigma, c0) = @@ -535,7 +535,7 @@ let pf_abs_evars2 gl rigid (sigma, c0) = let t = Context.Named.fold_inside abs_dc ~init:concl dc in nf_evar sigma t in let rec put evlist c = match Constr.kind c with - | Evar (k, a) -> + | Evar (k, a) -> if List.mem_assoc k evlist || Evd.mem sigma0 k || List.mem k rigid then evlist else let n = max 0 (Array.length a - nenv) in let t = abs_evar n k in (k, (n, t)) :: put evlist t @@ -561,11 +561,11 @@ let pf_abs_evars gl t = pf_abs_evars2 gl [] t (* As before but if (?i : T(?j)) and (?j : P : Prop), then the lambda for i - * looks like (fun evar_i : (forall pi : P. T(pi))) thanks to "loopP" and all + * looks like (fun evar_i : (forall pi : P. T(pi))) thanks to "loopP" and all * occurrences of evar_i are replaced by (evar_i evar_j) thanks to "app". * * If P can be solved by ssrautoprop (that defaults to trivial), then - * the corresponding lambda looks like (fun evar_i : T(c)) where c is + * the corresponding lambda looks like (fun evar_i : T(c)) where c is * the solution found by ssrautoprop. *) let ssrautoprop_tac = ref (fun gl -> assert false) @@ -596,11 +596,11 @@ let pf_abs_evars_pirrel gl (sigma, c0) = let t = Context.Named.fold_inside abs_dc ~init:concl dc in nf_evar sigma0 (nf_evar sigma t) in let rec put evlist c = match Constr.kind c with - | Evar (k, a) -> + | Evar (k, a) -> if List.mem_assoc k evlist || Evd.mem sigma0 k then evlist else let n = max 0 (Array.length a - nenv) in - let k_ty = - Retyping.get_sort_family_of + let k_ty = + Retyping.get_sort_family_of (pf_env gl) sigma (Evd.evar_concl (Evd.find sigma k)) in let is_prop = k_ty = InProp in let t = abs_evar n k in (k, (n, t, is_prop)) :: put evlist t @@ -610,23 +610,23 @@ let pf_abs_evars_pirrel gl (sigma, c0) = let pr_constr t = Printer.pr_econstr_env (pf_env gl) sigma (Reductionops.nf_beta (pf_env gl) (project gl) (EConstr.of_constr t)) in pp(lazy(str"evlist=" ++ pr_list (fun () -> str";") (fun (k,_) -> Evar.print k) evlist)); - let evplist = - let depev = List.fold_left (fun evs (_,(_,t,_)) -> + let evplist = + let depev = List.fold_left (fun evs (_,(_,t,_)) -> let t = EConstr.of_constr t in Intset.union evs (Evarutil.undefined_evars_of_term sigma t)) Intset.empty evlist in List.filter (fun (i,(_,_,b)) -> b && Intset.mem i depev) evlist in - let evlist, evplist, sigma = + let evlist, evplist, sigma = if evplist = [] then evlist, [], sigma else List.fold_left (fun (ev, evp, sigma) (i, (_,t,_) as p) -> - try + try let ng, sigma = call_on_evar !ssrautoprop_tac i sigma in if (ng <> []) then errorstrm (str "Should we tell the user?"); List.filter (fun (j,_) -> j <> i) ev, evp, sigma with _ -> ev, p::evp, sigma) (evlist, [], sigma) (List.rev evplist) in let c0 = nf_evar sigma c0 in - let evlist = + let evlist = List.map (fun (x,(y,t,z)) -> x,(y,nf_evar sigma t,z)) evlist in - let evplist = + let evplist = List.map (fun (x,(y,t,z)) -> x,(y,nf_evar sigma t,z)) evplist in pp(lazy(str"c0= " ++ pr_constr c0)); let rec lookup k i = function @@ -646,7 +646,7 @@ let pf_abs_evars_pirrel gl (sigma, c0) = let rec loopP evlist c i = function | (_, (n, t, _)) :: evl -> let t = get evlist (i - 1) t in - let n = Name (Id.of_string (ssr_anon_hyp ^ string_of_int n)) in + let n = Name (Id.of_string (ssr_anon_hyp ^ string_of_int n)) in loopP evlist (mkProd (make_annot n Sorts.Relevant, t, c)) (i - 1) evl | [] -> c in let rec loop c i = function @@ -655,8 +655,8 @@ let pf_abs_evars_pirrel gl (sigma, c0) = let t_evplist = List.filter (fun (k,_) -> Intset.mem k evs) evplist in let t = loopP t_evplist (get t_evplist 1 t) 1 t_evplist in let t = get evlist (i - 1) t in - let extra_args = - List.map (fun (k,_) -> mkRel (fst (lookup k i evlist))) + let extra_args = + List.map (fun (k,_) -> mkRel (fst (lookup k i evlist))) (List.rev t_evplist) in let c = if extra_args = [] then c else app extra_args 1 c in loop (mkLambda (make_annot (mk_evar_name n) Sorts.Relevant, t, c)) (i - 1) evl @@ -755,7 +755,7 @@ let pf_abs_prod name gl c cl = pf_mkprod gl c ~name (Termops.subst_term (project (** look up a name in the ssreflect internals module *) let ssrdirpath = DirPath.make [Id.of_string "ssreflect"] -let ssrqid name = Libnames.make_qualid ssrdirpath (Id.of_string name) +let ssrqid name = Libnames.make_qualid ssrdirpath (Id.of_string name) let mkSsrRef name = let qn = Format.sprintf "plugins.ssreflect.%s" name in if Coqlib.has_ref qn then Coqlib.lib_ref qn else @@ -858,7 +858,7 @@ let top_id = mk_internal_id "top assumption" let ssr_n_tac seed n gl = let name = if n = -1 then seed else ("ssr" ^ seed ^ string_of_int n) in let fail msg = CErrors.user_err (Pp.str msg) in - let tacname = + let tacname = try Tacenv.locate_tactic (Libnames.qualid_of_ident (Id.of_string name)) with Not_found -> try Tacenv.locate_tactic (ssrqid name) with Not_found -> @@ -927,13 +927,13 @@ let pf_interp_ty ?(resolve_typeclasses=false) ist gl ty = (* TASSI: given (c : ty), generates (c ??? : ty[???/...]) with m evars *) exception NotEnoughProducts -let saturate ?(beta=false) ?(bi_types=false) env sigma c ?(ty=Retyping.get_type_of env sigma c) m +let saturate ?(beta=false) ?(bi_types=false) env sigma c ?(ty=Retyping.get_type_of env sigma c) m = - let rec loop ty args sigma n = - if n = 0 then + let rec loop ty args sigma n = + if n = 0 then let args = List.rev args in (if beta then Reductionops.whd_beta sigma else fun x -> x) - (EConstr.mkApp (c, Array.of_list (List.map snd args))), ty, args, sigma + (EConstr.mkApp (c, Array.of_list (List.map snd args))), ty, args, sigma else match EConstr.kind_of_type sigma ty with | ProdType (_, src, tgt) -> let sigma = create_evar_defs sigma in @@ -941,7 +941,7 @@ let saturate ?(beta=false) ?(bi_types=false) env sigma c ?(ty=Retyping.get_type_ Evarutil.new_evar env sigma (if bi_types then Reductionops.nf_betaiota env sigma src else src) in loop (EConstr.Vars.subst1 x tgt) ((m - n,x) :: args) sigma (n-1) - | CastType (t, _) -> loop t args sigma n + | CastType (t, _) -> loop t args sigma n | LetInType (_, v, _, t) -> loop (EConstr.Vars.subst1 v t) args sigma n | SortType _ -> assert false | AtomicType _ -> @@ -953,10 +953,10 @@ let saturate ?(beta=false) ?(bi_types=false) env sigma c ?(ty=Retyping.get_type_ in loop ty [] sigma m -let pf_saturate ?beta ?bi_types gl c ?ty m = +let pf_saturate ?beta ?bi_types gl c ?ty m = let env, sigma, si = pf_env gl, project gl, sig_it gl in let t, ty, args, sigma = saturate ?beta ?bi_types env sigma c ?ty m in - t, ty, args, re_sig si sigma + t, ty, args, re_sig si sigma let pf_partial_solution gl t evl = let sigma, g = project gl, sig_it gl in @@ -973,7 +973,7 @@ let dependent_apply_error = * is just like apply, but with a user-provided number n of implicits. * * Refine.refine function that handles type classes and evars but fails to - * handle "dependently typed higher order evars". + * handle "dependently typed higher order evars". * * Refiner.refiner that does not handle metas with a non ground type but works * with dependently typed higher order metas. *) @@ -998,7 +998,7 @@ let applyn ~with_evars ?beta ?(with_shelve=false) ?(first_goes_last=false) n t g let t, gl = if n = 0 then t, gl else let sigma, si = project gl, sig_it gl in let rec loop sigma bo args = function (* saturate with metas *) - | 0 -> EConstr.mkApp (t, Array.of_list (List.rev args)), re_sig si sigma + | 0 -> EConstr.mkApp (t, Array.of_list (List.rev args)), re_sig si sigma | n -> match EConstr.kind sigma bo with | Lambda (_, ty, bo) -> if not (EConstr.Vars.closed0 sigma ty) then @@ -1041,7 +1041,7 @@ let introid ?(orig=ref Anonymous) name = tclTHEN (fun gl -> let g, env = Tacmach.pf_concl gl, pf_env gl in let sigma = project gl in match EConstr.kind sigma g with - | App (hd, _) when EConstr.isLambda sigma hd -> + | App (hd, _) when EConstr.isLambda sigma hd -> Proofview.V82.of_tactic (convert_concl_no_check (Reductionops.whd_beta sigma g)) gl | _ -> tclIDTAC gl) (Proofview.V82.of_tactic @@ -1066,9 +1066,9 @@ let is_pf_var sigma c = let hyp_of_var sigma v = SsrHyp (Loc.tag @@ EConstr.destVar sigma v) let interp_clr sigma = function -| Some clr, (k, c) +| Some clr, (k, c) when (k = xNoFlag || k = xWithAt) && is_pf_var sigma c -> - hyp_of_var sigma c :: clr + hyp_of_var sigma c :: clr | Some clr, _ -> clr | None, _ -> [] @@ -1091,7 +1091,7 @@ let tclDO n tac = let prefix i = str"At iteration " ++ int i ++ str": " in let tac_err_at i gl = try tac gl - with + with | CErrors.UserError (l, s) as e -> let _, info = CErrors.push e in let e' = CErrors.UserError (l, prefix i ++ s) in @@ -1123,7 +1123,7 @@ let pf_interp_gen_aux gl to_ind ((oclr, occ), t) = let pat = interp_cpattern gl t None in (* UGLY API *) let gl = pf_merge_uc_of (fst pat) gl in let cl, env, sigma = Tacmach.pf_concl gl, pf_env gl, project gl in - let (c, ucst), cl = + let (c, ucst), cl = try fill_occ_pattern ~raise_NoMatch:true env sigma (EConstr.Unsafe.to_constr cl) pat occ 1 with NoMatch -> redex_of_pattern env pat, (EConstr.Unsafe.to_constr cl) in let gl = pf_merge_uc ucst gl in @@ -1131,9 +1131,9 @@ let pf_interp_gen_aux gl to_ind ((oclr, occ), t) = let cl = EConstr.of_constr cl in let clr = interp_clr sigma (oclr, (tag_of_cpattern t, c)) in if not(occur_existential sigma c) then - if tag_of_cpattern t = xWithAt then + if tag_of_cpattern t = xWithAt then if not (EConstr.isVar sigma c) then - errorstrm (str "@ can be used with variables only") + errorstrm (str "@ can be used with variables only") else match Tacmach.pf_get_hyp gl (EConstr.destVar sigma c) with | NamedDecl.LocalAssum _ -> errorstrm (str "@ can be used with let-ins only") | NamedDecl.LocalDef (name, b, ty) -> true, pat, EConstr.mkLetIn (map_annot Name.mk_name name,b,ty,cl),c,clr,ucst,gl @@ -1186,7 +1186,7 @@ let gen_tmp_ids push_ctxs ctx (tclTHENLIST (List.map (fun (id,orig_ref) -> - tclTHEN + tclTHEN (gentac ((None,Some(false,[])),cpattern_of_id id)) (rename_hd_prod orig_ref)) ctx.tmp_ids) gl) @@ -1210,7 +1210,7 @@ let pfLIFT f = Proofview.tclUNIT x ;; -(* TASSI: This version of unprotects inlines the unfold tactic definition, +(* TASSI: This version of unprotects inlines the unfold tactic definition, * since we don't want to wipe out let-ins, and it seems there is no flag * to change that behaviour in the standard unfold code *) let unprotecttac gl = @@ -1219,8 +1219,8 @@ let unprotecttac gl = Tacticals.onClause (fun idopt -> let hyploc = Option.map (fun id -> id, InHyp) idopt in Proofview.V82.of_tactic (Tactics.reduct_option ~check:false - (Reductionops.clos_norm_flags - (CClosure.RedFlags.mkflags + (Reductionops.clos_norm_flags + (CClosure.RedFlags.mkflags [CClosure.RedFlags.fBETA; CClosure.RedFlags.fCONST prot; CClosure.RedFlags.fMATCH; @@ -1253,7 +1253,7 @@ let abs_wgen keep_let f gen (gl,args,c) = let x' = make_annot (Name (f x)) (NamedDecl.get_relevance hyp) in let prod = EConstr.mkProd (x', NamedDecl.get_type hyp, EConstr.Vars.subst_var x c) in gl, EConstr.mkVar x :: args, prod - | _, Some ((x, "@"), Some p) -> + | _, Some ((x, "@"), Some p) -> let x = hoi_id x in let cp = interp_cpattern gl p None in let gl = pf_merge_uc_of (fst cp) gl in diff --git a/plugins/ssr/ssrcommon.mli b/plugins/ssr/ssrcommon.mli index db1d2d456e..741db9a6c2 100644 --- a/plugins/ssr/ssrcommon.mli +++ b/plugins/ssr/ssrcommon.mli @@ -38,7 +38,7 @@ val hoi_id : ssrhyp_or_id -> Id.t (******************************* hints ***********************************) -val mk_hint : 'a -> 'a ssrhint +val mk_hint : 'a -> 'a ssrhint val mk_orhint : 'a -> bool * 'a val nullhint : bool * 'a list val nohint : 'a ssrhint @@ -122,7 +122,7 @@ val mkCLambda : ?loc:Loc.t -> Name.t -> constr_expr -> constr_expr -> constr_e val isCHoles : constr_expr list -> bool val isCxHoles : (constr_expr * 'a option) list -> bool -val intern_term : +val intern_term : Tacinterp.interp_sign -> env -> ssrterm -> Glob_term.glob_constr @@ -152,7 +152,7 @@ val pf_e_type_of : Goal.goal Evd.sigma -> EConstr.constr -> Goal.goal Evd.sigma * EConstr.types -val splay_open_constr : +val splay_open_constr : Goal.goal Evd.sigma -> evar_map * EConstr.t -> (Names.Name.t Context.binder_annot * EConstr.t) list * EConstr.t @@ -181,7 +181,7 @@ val mk_evar_name : int -> Name.t val ssr_anon_hyp : string val pf_type_id : Goal.goal Evd.sigma -> EConstr.types -> Id.t -val pf_abs_evars : +val pf_abs_evars : Goal.goal Evd.sigma -> evar_map * EConstr.t -> int * EConstr.t * Evar.t list * @@ -235,7 +235,7 @@ val is_discharged_id : Id.t -> bool val mk_discharged_id : Id.t -> Id.t val is_tagged : string -> string -> bool val has_discharged_tag : string -> bool -val ssrqid : string -> Libnames.qualid +val ssrqid : string -> Libnames.qualid val new_tmp_id : tac_ctx -> (Names.Id.t * Name.t ref) * tac_ctx val mk_anon_id : string -> Id.t list -> Id.t @@ -244,7 +244,7 @@ val pf_abs_evars_pirrel : evar_map * Constr.constr -> int * Constr.constr val nbargs_open_constr : Goal.goal Evd.sigma -> Evd.evar_map * EConstr.t -> int val pf_nbargs : Goal.goal Evd.sigma -> EConstr.t -> int -val gen_tmp_ids : +val gen_tmp_ids : ?ist:Geninterp.interp_sign -> (Goal.goal * tac_ctx) Evd.sigma -> (Goal.goal * tac_ctx) list Evd.sigma diff --git a/plugins/ssr/ssrelim.ml b/plugins/ssr/ssrelim.ml index d0426c86b9..26962ee87b 100644 --- a/plugins/ssr/ssrelim.ml +++ b/plugins/ssr/ssrelim.ml @@ -36,7 +36,7 @@ module RelDecl = Context.Rel.Declaration * checks if the eliminator is recursive or not *) let analyze_eliminator elimty env sigma = let rec loop ctx t = match EConstr.kind_of_type sigma t with - | AtomicType (hd, args) when EConstr.isRel sigma hd -> + | AtomicType (hd, args) when EConstr.isRel sigma hd -> ctx, EConstr.destRel sigma hd, not (EConstr.Vars.noccurn sigma 1 t), Array.length args, t | CastType (t, _) -> loop ctx t | ProdType (x, ty, t) -> loop (RelDecl.LocalAssum (x, ty) :: ctx) t @@ -50,7 +50,7 @@ let analyze_eliminator elimty env sigma = str"the eliminator's"++Pp.cut()++str"type:"++spc()++pr_econstr_env env' sigma elimty) in let ctx, pred_id, elim_is_dep, n_pred_args,concl = loop [] elimty in let n_elim_args = Context.Rel.nhyps ctx in - let is_rec_elim = + let is_rec_elim = let count_occurn n term = let count = ref 0 in let rec occur_rec n c = match EConstr.kind sigma c with @@ -59,7 +59,7 @@ let analyze_eliminator elimty env sigma = in occur_rec n term; !count in let occurr2 n t = count_occurn n t > 1 in - not (List.for_all_i + not (List.for_all_i (fun i (_,rd) -> pred_id <= i || not (occurr2 (pred_id - i) rd)) 1 (assums_of_rel_context ctx)) in @@ -68,7 +68,7 @@ let analyze_eliminator elimty env sigma = let subgoals_tys sigma (relctx, concl) = let rec aux cur_depth acc = function - | hd :: rest -> + | hd :: rest -> let ty = Context.Rel.Declaration.get_type hd in if EConstr.Vars.noccurn sigma cur_depth concl && List.for_all_i (fun i -> function @@ -94,7 +94,7 @@ let subgoals_tys sigma (relctx, concl) = * 1. find the eliminator if not given as ~elim and analyze it * 2. build the patterns to be matched against the conclusion, looking at * (occ, c), deps and the pattern inferred from the type of the eliminator - * 3. build the new predicate matching the patterns, and the tactic to + * 3. build the new predicate matching the patterns, and the tactic to * generalize the equality in case eqid is not None * 4. build the tactic handle instructions and clears as required in ipats and * by eqid *) @@ -131,7 +131,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = let is_undef_pat = function | sigma, T t -> EConstr.isEvar sigma (EConstr.of_constr t) | _ -> false in - let match_pat env p occ h cl = + let match_pat env p occ h cl = let sigma0 = project orig_gl in ppdebug(lazy Pp.(str"matching: " ++ pr_occ occ ++ pp_pattern env p)); let (c,ucst), cl = @@ -139,11 +139,11 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = ppdebug(lazy Pp.(str" got: " ++ pr_constr_env env sigma0 c)); c, EConstr.of_constr cl, ucst in let mkTpat gl t = (* takes a term, refreshes it and makes a T pattern *) - let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in + let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in let t, _, _, sigma = saturate ~beta:true env (project gl) t n in Evd.merge_universe_context sigma ucst, T (EConstr.Unsafe.to_constr t) in let unif_redex gl (sigma, r as p) t = (* t is a hint for the redex of p *) - let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in + let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in let t, _, _, sigma = saturate ~beta:true env sigma t n in let sigma = Evd.merge_universe_context sigma ucst in match r with @@ -317,11 +317,11 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = (* if we are the index for the equation we do not clear *) let clr_t = if deps = [] && eqid <> None then [] else clr_t in let p = if is_undef_pat p then mkTpat gl inf_t else p in - loop (patterns @ [i, p, inf_t, occ]) + loop (patterns @ [i, p, inf_t, occ]) (clr_t @ clr) (i+1) (deps, inf_deps) - | [], c :: inf_deps -> + | [], c :: inf_deps -> ppdebug(lazy Pp.(str"adding inf pattern " ++ pr_econstr_pat env (project gl) c)); - loop (patterns @ [i, mkTpat gl c, c, allocc]) + loop (patterns @ [i, mkTpat gl c, c, allocc]) clr (i+1) ([], inf_deps) | _::_, [] -> errorstrm Pp.(str "Too many dependent abstractions") in let deps, head_p, inf_deps_r = match what, c_is_head_p, cty with @@ -332,7 +332,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = let occ = if occ = None then allocc else occ in let inf_p, inf_deps_r = List.hd inf_deps_r, List.tl inf_deps_r in deps, [1, pc, inf_p, occ], inf_deps_r in - let patterns, clr, gl = + let patterns, clr, gl = loop [] orig_clr (List.length head_p+1) (List.rev deps, inf_deps_r) in head_p @ patterns, Util.List.uniquize clr, gl in @@ -340,7 +340,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = ppdebug(lazy Pp.(pp_concat (str"inf. patterns=") (List.map (pp_inf_pat gl) patterns))); (* Predicate generation, and (if necessary) tactic to generalize the * equation asked by the user *) - let elim_pred, gen_eq_tac, clr, gl = + let elim_pred, gen_eq_tac, clr, gl = let error gl t inf_t = errorstrm Pp.(str"The given pattern matches the term"++ spc()++pp_term gl t++spc()++str"while the inferred pattern"++ spc()++pr_econstr_pat env (project gl) (fire_subst gl inf_t)++spc()++ str"doesn't") in @@ -356,19 +356,19 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = let gl = try pf_unify_HO gl inf_t c with exn when CErrors.noncritical exn -> error gl c inf_t in cl, gl, post - with + with | NoMatch | NoProgress -> let e, ucst = redex_of_pattern env p in let gl = pf_merge_uc ucst gl in let e = EConstr.of_constr e in let n, e, _, _ucst = pf_abs_evars gl (fst p, e) in - let e, _, _, gl = pf_saturate ~beta:true gl e n in + let e, _, _, gl = pf_saturate ~beta:true gl e n in let gl = try pf_unify_HO gl inf_t e with exn when CErrors.noncritical exn -> error gl e inf_t in cl, gl, post - in + in let rec match_all concl gl patterns = - let concl, gl, postponed = + let concl, gl, postponed = List.fold_left match_or_postpone (concl, gl, []) patterns in if postponed = [] then concl, gl else if List.length postponed = List.length patterns then @@ -377,8 +377,8 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = else match_all concl gl postponed in let concl, gl = match_all concl gl patterns in let pred_rctx, _ = EConstr.decompose_prod_assum (project gl) (fire_subst gl predty) in - let concl, gen_eq_tac, clr, gl = match eqid with - | Some (IPatId _) when not is_rec -> + let concl, gen_eq_tac, clr, gl = match eqid with + | Some (IPatId _) when not is_rec -> let k = List.length deps in let c = fire_subst gl (List.assoc (n_elim_args - k - 1) elim_args) in let gl, t = pfe_type_of gl c in @@ -405,7 +405,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = | _ -> concl, Tacticals.tclIDTAC, clr, gl in let mk_lam t r = EConstr.mkLambda_or_LetIn r t in let concl = List.fold_left mk_lam concl pred_rctx in - let gl, concl = + let gl, concl = if eqid <> None && is_rec then let gl, concls = pfe_type_of gl concl in let concl, gl = mkProt concls concl gl in @@ -421,10 +421,10 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = let gl = pf_resolve_typeclasses ~where:elim ~fail:false gl in let gl, _ = pf_e_type_of gl elim in (* check that the patterns do not contain non instantiated dependent metas *) - let () = + let () = let evars_of_term = Evarutil.undefined_evars_of_term (project gl) in let patterns = List.map (fun (_,_,t,_) -> fire_subst gl t) patterns in - let patterns_ev = List.map evars_of_term patterns in + let patterns_ev = List.map evars_of_term patterns in let ev = List.fold_left Evar.Set.union Evar.Set.empty patterns_ev in let ty_ev = Evar.Set.fold (fun i e -> let ex = i in @@ -441,7 +441,7 @@ let ssrelim ?(is_case=false) deps what ?elim eqid elim_intro_tac = end in (* the elim tactic, with the eliminator and the predicated we computed *) - let elim = project gl, elim in + let elim = project gl, elim in let seed = Array.map (fun ty -> let ctx,_ = EConstr.decompose_prod_assum (project gl) ty in @@ -517,7 +517,7 @@ let perform_injection c gl = let cl1 = EConstr.mkLambda EConstr.(make_annot Anonymous Sorts.Relevant, mkArrow eqt Sorts.Relevant cl, mkApp (mkRel 1, [|c_eq|])) in let id = injecteq_id in let id_with_ebind = (EConstr.mkVar id, NoBindings) in - let injtac = Tacticals.tclTHEN (introid id) (injectidl2rtac id id_with_ebind) in + let injtac = Tacticals.tclTHEN (introid id) (injectidl2rtac id id_with_ebind) in Tacticals.tclTHENLAST (Proofview.V82.of_tactic (Tactics.apply (EConstr.compose_lam dc cl1))) injtac gl let ssrscase_or_inj_tac c = Proofview.V82.tactic ~nf_evars:false (fun gl -> diff --git a/plugins/ssr/ssrequality.ml b/plugins/ssr/ssrequality.ml index 742890637a..cdda84a18d 100644 --- a/plugins/ssr/ssrequality.ml +++ b/plugins/ssr/ssrequality.ml @@ -125,8 +125,8 @@ let newssrcongrtac arg ist gl = | Some gl_c -> tclTHEN (Proofview.V82.of_tactic (convert_concl ~check:true (fs gl_c c))) (t_ok (proj gl_c)) gl - | None -> t_fail () gl in - let mk_evar gl ty = + | None -> t_fail () gl in + let mk_evar gl ty = let env, sigma, si = pf_env gl, project gl, sig_it gl in let sigma = Evd.create_evar_defs sigma in let (sigma, x) = Evarutil.new_evar env sigma ty in @@ -174,7 +174,7 @@ let nodocc = mkclr [] let is_rw_cut = function RWred (Cut _) -> true | _ -> false -let mk_rwarg (d, (n, _ as m)) ((clr, occ as docc), rx) (rt, _ as r) : ssrrwarg = +let mk_rwarg (d, (n, _ as m)) ((clr, occ as docc), rx) (rt, _ as r) : ssrrwarg = if rt <> RWeq then begin if rt = RWred Nop && not (m = nomult && occ = None && rx = None) && (clr = None || clr = Some []) then @@ -190,7 +190,7 @@ let mk_rwarg (d, (n, _ as m)) ((clr, occ as docc), rx) (rt, _ as r) : ssrrwarg = let norwmult = L2R, nomult let norwocc = noclr, None -let simplintac occ rdx sim gl = +let simplintac occ rdx sim gl = let simptac m gl = if m <> ~-1 then begin if rdx <> None then @@ -219,7 +219,7 @@ let rec get_evalref env sigma c = match EConstr.kind sigma c with (* Strip a pattern generated by a prenex implicit to its constant. *) let strip_unfold_term _ ((sigma, t) as p) kt = match EConstr.kind sigma t with - | App (f, a) when kt = xNoFlag && Array.for_all (EConstr.isEvar sigma) a && EConstr.isConst sigma f -> + | App (f, a) when kt = xNoFlag && Array.for_all (EConstr.isEvar sigma) a && EConstr.isConst sigma f -> (sigma, f), true | Const _ | Var _ -> p, true | Proj _ -> p, true @@ -235,7 +235,7 @@ let all_ok _ _ = true let fake_pmatcher_end () = mkProp, L2R, (Evd.empty, UState.empty, mkProp) -let unfoldintac occ rdx t (kt,_) gl = +let unfoldintac occ rdx t (kt,_) gl = let fs sigma x = Reductionops.nf_evar sigma x in let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in let (sigma, t), const = strip_unfold_term env0 t kt in @@ -250,18 +250,18 @@ let unfoldintac occ rdx t (kt,_) gl = let ise, u = mk_tpattern env0 sigma0 (ise,EConstr.Unsafe.to_constr t) all_ok L2R (EConstr.Unsafe.to_constr t) in let find_T, end_T = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in - (fun env c _ h -> + (fun env c _ h -> try find_T env c h ~k:(fun env c _ _ -> EConstr.Unsafe.to_constr (body env t (EConstr.of_constr c))) with NoMatch when easy -> c | NoMatch | NoProgress -> errorstrm Pp.(str"No occurrence of " ++ pr_econstr_pat env sigma0 t ++ spc() ++ str "in " ++ Printer.pr_constr_env env sigma c)), - (fun () -> try end_T () with - | NoMatch when easy -> fake_pmatcher_end () + (fun () -> try end_T () with + | NoMatch when easy -> fake_pmatcher_end () | NoMatch -> anomaly "unfoldintac") - | _ -> + | _ -> (fun env (c as orig_c) _ h -> if const then - let rec aux c = + let rec aux c = match EConstr.kind sigma0 c with | Const _ when EConstr.eq_constr sigma0 c t -> body env t t | App (f,a) when EConstr.eq_constr sigma0 f t -> EConstr.mkApp (body env f f,a) @@ -282,15 +282,15 @@ let unfoldintac occ rdx t (kt,_) gl = with _ -> errorstrm Pp.(str "The term " ++ pr_constr_env env sigma c ++spc()++ str "does not unify with " ++ pr_econstr_pat env sigma t)), fake_pmatcher_end in - let concl = + let concl = let concl0 = EConstr.Unsafe.to_constr concl0 in - try beta env0 (EConstr.of_constr (eval_pattern env0 sigma0 concl0 rdx occ unfold)) + try beta env0 (EConstr.of_constr (eval_pattern env0 sigma0 concl0 rdx occ unfold)) with Option.IsNone -> errorstrm Pp.(str"Failed to unfold " ++ pr_econstr_pat env0 sigma t) in let _ = conclude () in Proofview.V82.of_tactic (convert_concl ~check:true concl) gl ;; -let foldtac occ rdx ft gl = +let foldtac occ rdx ft gl = let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in let sigma, t = ft in let t = EConstr.to_constr ~abort_on_undefined_evars:false sigma t in @@ -303,7 +303,7 @@ let foldtac occ rdx ft gl = mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[ut]) in (fun env c _ h -> try find_T env c h ~k:(fun env t _ _ -> t) with NoMatch ->c), (fun () -> try end_T () with NoMatch -> fake_pmatcher_end ()) - | _ -> + | _ -> (fun env c _ h -> try let sigma = unify_HO env sigma (EConstr.of_constr c) (EConstr.of_constr t) in @@ -371,12 +371,12 @@ let pirrel_rewrite ?(under=false) ?(map_redex=id_map_redex) pred rdx rdx_ty new_ in ppdebug(lazy Pp.(str"pirrel_rewrite: proof term: " ++ pr_econstr_env env sigma proof)); ppdebug(lazy Pp.(str"pirrel_rewrite of type: " ++ pr_econstr_env env sigma proof_ty)); - try refine_with + try refine_with ~first_goes_last:(not !ssroldreworder || under) ~with_evars:under (sigma, proof) gl - with _ -> + with _ -> (* we generate a msg like: "Unable to find an instance for the variable" *) let hd_ty, miss = match EConstr.kind sigma c with - | App (hd, args) -> + | App (hd, args) -> let hd_ty = Retyping.get_type_of env sigma hd in let names = let rec aux t = function 0 -> [] | n -> let t = Reductionops.whd_all env sigma t in @@ -409,7 +409,7 @@ let rwcltac ?under ?map_redex cl rdx dir sr gl = (* ppdebug(lazy(str"sigma@rwcltac=" ++ pr_evar_map None (fst sr))); *) ppdebug(lazy Pp.(str"r@rwcltac=" ++ pr_econstr_env (pf_env gl) (project gl) (snd sr))); let cvtac, rwtac, gl = - if EConstr.Vars.closed0 (project gl) r' then + if EConstr.Vars.closed0 (project gl) r' then let env, sigma, c, c_eq = pf_env gl, fst sr, snd sr, Coqlib.(lib_ref "core.eq.type") in let sigma, c_ty = Typing.type_of env sigma c in ppdebug(lazy Pp.(str"c_ty@rwcltac=" ++ pr_econstr_env env sigma c_ty)); @@ -417,14 +417,14 @@ let rwcltac ?under ?map_redex cl rdx dir sr gl = | AtomicType(e, a) when Ssrcommon.is_ind_ref sigma e c_eq -> let new_rdx = if dir = L2R then a.(2) else a.(1) in pirrel_rewrite ?under ?map_redex cl rdx rdxt new_rdx dir (sigma,c) c_ty, tclIDTAC, gl - | _ -> + | _ -> let cl' = EConstr.mkApp (EConstr.mkNamedLambda (make_annot pattern_id Sorts.Relevant) rdxt cl, [|rdx|]) in let sigma, _ = Typing.type_of env sigma cl' in let gl = pf_merge_uc_of sigma gl in Proofview.V82.of_tactic (convert_concl ~check:true cl'), rewritetac ?under dir r', gl else let dc, r2 = EConstr.decompose_lam_n_assum (project gl) n r' in - let r3, _, r3t = + let r3, _, r3t = try EConstr.destCast (project gl) r2 with _ -> errorstrm Pp.(str "no cast from " ++ pr_econstr_pat (pf_env gl) (project gl) (snd sr) ++ str " to " ++ pr_econstr_env (pf_env gl) (project gl) r2) in @@ -471,7 +471,7 @@ let ssr_is_setoid env = | None -> fun _ _ _ -> false | Some srel -> fun sigma r args -> - Rewrite.is_applied_rewrite_relation env + Rewrite.is_applied_rewrite_relation env sigma [] (EConstr.mkApp (r, args)) <> None let closed0_check cl p gl = @@ -585,7 +585,7 @@ let rwrxtac ?under ?map_redex occ rdx_pat dir rule gl = sigma, pats @ [pat] in let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in let find_R, end_R = mk_tpattern_matcher sigma0 occ ~upats_origin rpats in - (fun e c _ i -> find_R ~k:(fun _ _ _ h -> mkRel h) e c i), + (fun e c _ i -> find_R ~k:(fun _ _ _ h -> mkRel h) e c i), fun cl -> let rdx,d,r = end_R () in closed0_check cl rdx gl; (d,r),rdx | Some(_, (T e | X_In_T (_,e) | E_As_X_In_T (e,_,_) | E_In_X_In_T (e,_,_))) -> let r = ref None in @@ -633,7 +633,7 @@ let rwargtac ?under ?map_redex ist ((dir, mult), (((oclr, occ), grx), (kind, gt) let interp gc gl = try interp_term ist gl gc with _ when snd mult = May -> fail := true; (project gl, EConstr.mkProp) in - let rwtac gl = + let rwtac gl = let rx = Option.map (interp_rpattern gl) grx in let gl = match rx with | None -> gl @@ -672,6 +672,6 @@ let unlocktac ist args gl = let locked, gl = pf_mkSsrConst "locked" gl in let key, gl = pf_mkSsrConst "master_key" gl in let ktacs = [ - (fun gl -> unfoldtac None None (project gl,locked) xInParens gl); + (fun gl -> unfoldtac None None (project gl,locked) xInParens gl); Proofview.V82.of_tactic (Ssrelim.casetac key (fun ?seed:_ k -> k)) ] in tclTHENLIST (List.map utac args @ ktacs) gl diff --git a/plugins/ssr/ssrfwd.ml b/plugins/ssr/ssrfwd.ml index b0f56c423f..f486d1e457 100644 --- a/plugins/ssr/ssrfwd.ml +++ b/plugins/ssr/ssrfwd.ml @@ -42,7 +42,7 @@ let ssrsettac id ((_, (pat, pty)), (_, occ)) gl = (mkRHole, Some body), ist) pty in let pat = interp_cpattern gl pat pty in let cl, sigma, env = pf_concl gl, project gl, pf_env gl in - let (c, ucst), cl = + let (c, ucst), cl = let cl = EConstr.Unsafe.to_constr cl in try fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1 with NoMatch -> redex_of_pattern ~resolve_typeclasses:true env pat, cl in @@ -77,8 +77,8 @@ let () = }) -open Constrexpr -open Glob_term +open Constrexpr +open Glob_term let combineCG t1 t2 f g = match t1, t2 with | (x, (t1, None)), (_, (t2, None)) -> x, (g t1 t2, None) @@ -96,7 +96,7 @@ let introstac ipats = Proofview.V82.of_tactic (tclIPAT ipats) let havetac ist (transp,((((clr, orig_pats), binders), simpl), (((fk, _), t), hint))) - suff namefst gl + suff namefst gl = let concl = pf_concl gl in let pats = tclCompileIPats orig_pats in @@ -195,7 +195,7 @@ let havetac ist | _,false,true -> let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in gl, EConstr.mkArrow ty Sorts.Relevant concl, hint, id, itac_c - | _, false, false -> + | _, false, false -> let n, cty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in gl, cty, Tacticals.tclTHEN (binderstac n) hint, id, Tacticals.tclTHEN itac_c simpltac | _, true, false -> assert false in @@ -260,7 +260,7 @@ let wlogtac ist (((clr0, pats),_),_) (gens, ((_, ct))) hint suff ghave gl = | _ -> CErrors.anomaly(str"SSR: wlog: pired: " ++ pr_econstr_env env sigma c) in c, args, pired c args, pf_merge_uc uc gl in let tacipat pats = introstac pats in - let tacigens = + let tacigens = Tacticals.tclTHEN (Tacticals.tclTHENLIST(List.rev(List.fold_right mkclr gens [old_cleartac clr0]))) (introstac (List.fold_right mkpats gens [])) in diff --git a/plugins/ssr/ssrparser.mlg b/plugins/ssr/ssrparser.mlg index a1f707ffa8..22325f3fc3 100644 --- a/plugins/ssr/ssrparser.mlg +++ b/plugins/ssr/ssrparser.mlg @@ -167,7 +167,7 @@ let pr_name = function Name id -> pr_id id | Anonymous -> str "_" let pr_spc () = str " " let pr_list = prlist_with_sep -(**************************** ssrhyp **************************************) +(**************************** ssrhyp **************************************) let pr_ssrhyp _ _ _ = pr_hyp @@ -279,7 +279,7 @@ let negate_parser f tok x = let rc = try Some (f tok x) with Stream.Failure -> None in match rc with | None -> () - | Some _ -> raise Stream.Failure + | Some _ -> raise Stream.Failure let test_not_ssrslashnum = Pcoq.Entry.of_parser @@ -612,10 +612,10 @@ let ipat_of_intro_pattern p = Tactypes.( | IntroAction IntroWildcard -> IPatAnon Drop | IntroAction (IntroOrAndPattern (IntroOrPattern iorpat)) -> IPatCase (Regular( - List.map (List.map ipat_of_intro_pattern) - (List.map (List.map remove_loc) iorpat))) + List.map (List.map ipat_of_intro_pattern) + (List.map (List.map remove_loc) iorpat))) | IntroAction (IntroOrAndPattern (IntroAndPattern iandpat)) -> - IPatCase + IPatCase (Regular [List.map ipat_of_intro_pattern (List.map remove_loc iandpat)]) | IntroNaming IntroAnonymous -> IPatAnon (One None) | IntroAction (IntroRewrite b) -> IPatRewrite (allocc, if b then L2R else R2L) @@ -688,7 +688,7 @@ let rec add_intro_pattern_hyps ipat hyps = | IntroAction (IntroRewrite _) -> hyps | IntroAction (IntroInjection ips) -> List.fold_right add_intro_pattern_hyps ips hyps | IntroAction (IntroApplyOn (c,pat)) -> add_intro_pattern_hyps pat hyps - | IntroForthcoming _ -> + | IntroForthcoming _ -> (* As in ipat_of_intro_pattern, was unable to determine which kind of ipat interp_introid could return [HH] *) assert false @@ -890,12 +890,12 @@ let check_ssrhpats loc w_binders ipats = let err_loc s = CErrors.user_err ~loc ~hdr:"ssreflect" s in let clr, ipats = let opt_app = function None -> fun l -> Some l - | Some l1 -> fun l2 -> Some (l1 @ l2) in + | Some l1 -> fun l2 -> Some (l1 @ l2) in let rec aux clr = function | IPatClear cl :: tl -> aux (opt_app clr cl) tl | tl -> clr, tl in aux None ipats in - let simpl, ipats = + let simpl, ipats = match List.rev ipats with | IPatSimpl _ as s :: tl -> [s], List.rev tl | _ -> [], ipats in @@ -906,7 +906,7 @@ let check_ssrhpats loc w_binders ipats = | [] -> ipat, [] | ( IPatId _| IPatAnon _| IPatCase _ | IPatDispatch _ | IPatRewrite _ as i) :: tl -> if w_binders then - if simpl <> [] && tl <> [] then + if simpl <> [] && tl <> [] then err_loc(str"binders XOR s-item allowed here: "++pr_ipats(tl@simpl)) else if not (List.for_all (function IPatId _ -> true | _ -> false) tl) then err_loc (str "Only binders allowed here: " ++ pr_ipats tl) @@ -939,7 +939,7 @@ ARGUMENT EXTEND ssrhpats_wtransp | [ ssripats(i) "@" ssripats(j) ] -> { true,check_ssrhpats loc true (i @ j) } END -ARGUMENT EXTEND ssrhpats_nobs +ARGUMENT EXTEND ssrhpats_nobs TYPED AS (((ssrclear option * ssripats) * ssripats) * ssripats) PRINTED BY { pr_ssrhpats } | [ ssripats(i) ] -> { check_ssrhpats loc false i } END @@ -1019,7 +1019,7 @@ GRAMMAR EXTEND Gram ssrfwdid: [[ test_ssrfwdid; id = Prim.ident -> { id } ]]; END - + (* by *) (** Tactical arguments. *) @@ -1109,7 +1109,7 @@ END open Ssrmatching_plugin.Ssrmatching open Ssrmatching_plugin.G_ssrmatching -let pr_wgen = function +let pr_wgen = function | (clr, Some((id,k),None)) -> spc() ++ pr_clear mt clr ++ str k ++ pr_hoi id | (clr, Some((id,k),Some p)) -> spc() ++ pr_clear mt clr ++ str"(" ++ str k ++ pr_hoi id ++ str ":=" ++ @@ -1152,7 +1152,7 @@ let pr_ssrclausehyps _ _ _ = pr_clausehyps } -ARGUMENT EXTEND ssrclausehyps +ARGUMENT EXTEND ssrclausehyps TYPED AS ssrwgen list PRINTED BY { pr_ssrclausehyps } | [ ssrwgen(hyp) "," ssrclausehyps(hyps) ] -> { hyp :: hyps } | [ ssrwgen(hyp) ssrclausehyps(hyps) ] -> { hyp :: hyps } @@ -1163,7 +1163,7 @@ END (* type ssrclauses = ssrahyps * ssrclseq *) -let pr_clauses (hyps, clseq) = +let pr_clauses (hyps, clseq) = if clseq = InGoal then mt () else str "in " ++ pr_clausehyps hyps ++ pr_clseq clseq let pr_ssrclauses _ _ _ = pr_clauses @@ -1215,7 +1215,7 @@ let rec format_local_binders h0 bl0 = match h0, bl0 with | BFdef :: h, CLocalDef ({CAst.v=x}, v, oty) :: bl -> Bdef (x, oty, v) :: format_local_binders h bl | _ -> [] - + let rec format_constr_expr h0 c0 = let open CAst in match h0, c0 with | BFvar :: h, { v = CLambdaN ([CLocalAssum([{CAst.v=x}], _, _)], c) } -> let bs, c' = format_constr_expr h c in @@ -1228,11 +1228,11 @@ let rec format_constr_expr h0 c0 = let open CAst in match h0, c0 with Bdef (x, oty, v) :: bs, c' | [BFcast], { v = CCast (c, Glob_term.CastConv t) } -> [Bcast t], c - | BFrec (has_str, has_cast) :: h, + | BFrec (has_str, has_cast) :: h, { v = CFix ( _, [_, Some {CAst.v = CStructRec locn}, bl, t, c]) } -> let bs = format_local_binders h bl in let bstr = if has_str then [Bstruct (Name locn.CAst.v)] else [] in - bs @ bstr @ (if has_cast then [Bcast t] else []), c + bs @ bstr @ (if has_cast then [Bcast t] else []), c | BFrec (_, has_cast) :: h, { v = CCoFix ( _, [_, bl, t, c]) } -> format_local_binders h bl @ (if has_cast then [Bcast t] else []), c | _, c -> @@ -1516,7 +1516,7 @@ END { -let intro_id_to_binder = List.map (function +let intro_id_to_binder = List.map (function | IPatId id -> let { CAst.loc=xloc } as x = bvar_lname (mkCVar id) in (FwdPose, [BFvar]), @@ -1687,7 +1687,7 @@ let ssr_null_entry = Pcoq.Entry.of_parser "ssr_null" (fun _ _ -> ()) } -GRAMMAR EXTEND Gram +GRAMMAR EXTEND Gram GLOBAL: Prim.ident; Prim.ident: [[ s = IDENT; ssr_null_entry -> { ssr_id_of_string loc s } ]]; END @@ -1756,8 +1756,8 @@ END { let ssrautoprop gl = - try - let tacname = + try + let tacname = try Tacenv.locate_tactic (qualid_of_ident (Id.of_string "ssrautoprop")) with Not_found -> Tacenv.locate_tactic (ssrqid "ssrautoprop") in let tacexpr = CAst.make @@ Tacexpr.Reference (ArgArg (Loc.tag @@ tacname)) in @@ -2168,7 +2168,7 @@ let pr_ssraarg _ _ _ (view, (dgens, ipats)) = } -ARGUMENT EXTEND ssrapplyarg +ARGUMENT EXTEND ssrapplyarg TYPED AS (ssrbwdview * (ssragens * ssrintros)) PRINTED BY { pr_ssraarg } | [ ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> @@ -2567,7 +2567,7 @@ ARGUMENT EXTEND ssrwlogfwd TYPED AS (ssrwgen list * ssrfwd) | [ ":" ssrwgen_list(gens) "/" ast_closure_lterm(t) ] -> { gens, mkFwdHint "/" t} END - + TACTIC EXTEND ssrwlog | [ "wlog" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> { V82.tactic (wlogtac ist pats fwd hint false `NoGen) } @@ -2589,13 +2589,13 @@ TACTIC EXTEND ssrwithoutloss END TACTIC EXTEND ssrwithoutlosss -| [ "without" "loss" "suff" +| [ "without" "loss" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } END TACTIC EXTEND ssrwithoutlossss -| [ "without" "loss" "suffices" +| [ "without" "loss" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> { V82.tactic (wlogtac ist pats fwd hint true `NoGen) } END @@ -2627,7 +2627,7 @@ let test_idcomma = Pcoq.Entry.of_parser "test_idcomma" accept_idcomma GRAMMAR EXTEND Gram GLOBAL: ssr_idcomma; - ssr_idcomma: [ [ test_idcomma; + ssr_idcomma: [ [ test_idcomma; ip = [ id = IDENT -> { Some (Id.of_string id) } | "_" -> { None } ]; "," -> { Some ip } ] ]; diff --git a/plugins/ssr/ssrprinters.mli b/plugins/ssr/ssrprinters.mli index e4df7399e1..240b1a5476 100644 --- a/plugins/ssr/ssrprinters.mli +++ b/plugins/ssr/ssrprinters.mli @@ -17,7 +17,7 @@ val pp_term : val pr_spc : unit -> Pp.t val pr_bar : unit -> Pp.t -val pr_list : +val pr_list : (unit -> Pp.t) -> ('a -> Pp.t) -> 'a list -> Pp.t val pp_concat : diff --git a/plugins/ssr/ssrtacticals.ml b/plugins/ssr/ssrtacticals.ml index cd2448d764..0fc05f58d3 100644 --- a/plugins/ssr/ssrtacticals.ml +++ b/plugins/ssr/ssrtacticals.ml @@ -109,7 +109,7 @@ let endclausestac id_map clseq gl_id cl0 gl = EConstr.mkLetIn ({na with binder_name=Name (orig_id id)}, unmark v, unmark t, unmark c') | _ -> EConstr.map (project gl) unmark c in let utac hyp = - Proofview.V82.of_tactic + Proofview.V82.of_tactic (Tactics.convert_hyp ~check:false ~reorder:false (NamedDecl.map_constr unmark hyp)) in let utacs = List.map utac (pf_hyps gl) in let ugtac gl' = diff --git a/plugins/ssr/ssrvernac.mlg b/plugins/ssr/ssrvernac.mlg index 064ea0a3e3..9f6fe0e651 100644 --- a/plugins/ssr/ssrvernac.mlg +++ b/plugins/ssr/ssrvernac.mlg @@ -104,7 +104,7 @@ GRAMMAR EXTEND Gram [ "if"; c = operconstr LEVEL "200"; "is"; db1 = ssr_dthen; b2 = ssr_else -> { let b1, ct, rt = db1 in CAst.make ~loc @@ CCases (MatchStyle, rt, [mk_pat c ct], [b1; b2]) } | "if"; c = operconstr LEVEL "200";"isn't";db1 = ssr_dthen; b2 = ssr_else -> - { let b1, ct, rt = db1 in + { let b1, ct, rt = db1 in let b1, b2 = let open CAst in let {loc=l1; v=(p1, r1)}, {loc=l2; v=(p2, r2)} = b1, b2 in (make ?loc:l1 (p1, r2), make ?loc:l2 (p2, r1)) @@ -147,7 +147,7 @@ END let declare_one_prenex_implicit locality f = let fref = - try Smartlocate.global_with_alias f + try Smartlocate.global_with_alias f with _ -> errorstrm (pr_qualid f ++ str " is not declared") in let rec loop = function | a :: args' when Impargs.is_status_implicit a -> @@ -340,7 +340,7 @@ END (* Main type conclusion pattern filter *) -let rec splay_search_pattern na = function +let rec splay_search_pattern na = function | Pattern.PApp (fp, args) -> splay_search_pattern (na + Array.length args) fp | Pattern.PLetIn (_, _, _, bp) -> splay_search_pattern na bp | Pattern.PRef hr -> hr, na @@ -364,11 +364,11 @@ let coerce_search_pattern_to_sort hpat = if np < na then CErrors.user_err (Pp.str "too many arguments in head search pattern") else let hpat' = if np = na then hpat else mkPApp hpat (np - na) [||] in let warn () = - Feedback.msg_warning (str "Listing only lemmas with conclusion matching " ++ + Feedback.msg_warning (str "Listing only lemmas with conclusion matching " ++ pr_constr_pattern_env env sigma hpat') in if EConstr.isSort sigma ht then begin warn (); true, hpat' end else let filter_head, coe_path = - try + try let _, cp = Classops.lookup_path_to_sort_from (push_rels_assum dc env) sigma ht in warn (); diff --git a/plugins/ssrmatching/ssrmatching.ml b/plugins/ssrmatching/ssrmatching.ml index 9682487a22..6cb464918a 100644 --- a/plugins/ssrmatching/ssrmatching.ml +++ b/plugins/ssrmatching/ssrmatching.ml @@ -310,7 +310,7 @@ let pf_unify_HO gl t1 t2 = (* This is what the definition of iter_constr should be... *) let iter_constr_LR f c = match kind c with | Evar (k, a) -> Array.iter f a - | Cast (cc, _, t) -> f cc; f t + | Cast (cc, _, t) -> f cc; f t | Prod (_, t, b) | Lambda (_, t, b) -> f t; f b | LetIn (_, v, t, b) -> f v; f t; f b | App (cf, a) -> f cf; Array.iter f a @@ -423,10 +423,10 @@ let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p = | Var _ | Ind _ | Construct _ -> KpatFixed, f, a | Evar (k, _) -> if Evd.mem sigma0 k then KpatEvar k, f, a else - if a <> [] then KpatFlex, f, a else + if a <> [] then KpatFlex, f, a else (match p_origin with None -> CErrors.user_err Pp.(str "indeterminate pattern") | Some (dir, rule) -> - errorstrm (str "indeterminate " ++ pr_dir_side dir + errorstrm (str "indeterminate " ++ pr_dir_side dir ++ str " in " ++ pr_constr_pat env ise rule)) | LetIn (_, v, _, b) -> if b <> mkRel 1 then KpatLet, f, a else KpatFlex, v, a @@ -435,7 +435,7 @@ let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p = let aa = Array.of_list a in let ise', p' = evars_for_FO ~hack env sigma0 ise (mkApp (f, aa)) in ise', - { up_k = k; up_FO = p'; up_f = f; + { up_k = k; up_FO = p'; up_f = f; up_a = aa; up_ok = ok; up_dir = dir; up_t = t} (* Specialize a pattern after a successful match: assign a precise head *) @@ -462,7 +462,7 @@ let nb_cs_proj_args pc f u = try match kind f with | Prod _ -> na Prod_cs | Sort s -> na (Sort_cs (Sorts.family s)) - | Const (c',_) when Constant.equal c' pc -> nargs_of_proj u.up_f + | Const (c',_) when Constant.equal c' pc -> nargs_of_proj u.up_f | Proj (c',_) when Constant.equal (Projection.constant c') pc -> nargs_of_proj u.up_f | Var _ | Ind _ | Construct _ | Const _ -> na (Const_cs (global_of_constr f)) | _ -> -1 @@ -636,15 +636,15 @@ let match_upats_HO ~on_instance upats env sigma0 ise c = let fixed_upat evd = function -| {up_k = KpatFlex | KpatEvar _ | KpatProj _} -> false +| {up_k = KpatFlex | KpatEvar _ | KpatProj _} -> false | {up_t = t} -> not (occur_existential evd (EConstr.of_constr t)) (** FIXME *) let do_once r f = match !r with Some _ -> () | None -> r := Some (f ()) -let assert_done r = +let assert_done r = match !r with Some x -> x | None -> CErrors.anomaly (str"do_once never called.") -let assert_done_multires r = +let assert_done_multires r = match !r with | None -> CErrors.anomaly (str"do_once never called.") | Some (e, n, xs) -> @@ -652,7 +652,7 @@ let assert_done_multires r = try List.nth xs n with Failure _ -> raise NoMatch type subst = Environ.env -> constr -> constr -> int -> constr -type find_P = +type find_P = Environ.env -> constr -> int -> k:subst -> constr @@ -677,7 +677,7 @@ let mk_tpattern_matcher ?(all_instances=false) if !nocc = max_occ then skip_occ := use_occ; if !nocc <= max_occ then occ_set.(!nocc - 1) else not use_occ in let upat_that_matched = ref None in - let match_EQ env sigma u = + let match_EQ env sigma u = match u.up_k with | KpatLet -> let x, pv, t, pb = destLetIn u.up_f in @@ -693,14 +693,14 @@ let mk_tpattern_matcher ?(all_instances=false) | Lambda _ -> unif_EQ env sigma u.up_f c | _ -> false) | _ -> unif_EQ env sigma u.up_f in -let p2t p = mkApp(p.up_f,p.up_a) in +let p2t p = mkApp(p.up_f,p.up_a) in let source env = match upats_origin, upats with - | None, [p] -> + | None, [p] -> (if fixed_upat ise p then str"term " else str"partial term ") ++ pr_constr_pat env ise (p2t p) ++ spc() - | Some (dir,rule), [p] -> str"The " ++ pr_dir_side dir ++ str" of " ++ + | Some (dir,rule), [p] -> str"The " ++ pr_dir_side dir ++ str" of " ++ pr_constr_pat env ise rule ++ fnl() ++ ws 4 ++ pr_constr_pat env ise (p2t p) ++ fnl() - | Some (dir,rule), _ -> str"The " ++ pr_dir_side dir ++ str" of " ++ + | Some (dir,rule), _ -> str"The " ++ pr_dir_side dir ++ str" of " ++ pr_constr_pat env ise rule ++ spc() | _, [] | None, _::_::_ -> CErrors.anomaly (str"mk_tpattern_matcher with no upats_origin.") in @@ -724,8 +724,8 @@ let rec uniquize = function equal f f1 && CArray.for_all2 equal a a1) in x :: uniquize (List.filter neq xs) in -((fun env c h ~k -> - do_once upat_that_matched (fun () -> +((fun env c h ~k -> + do_once upat_that_matched (fun () -> let failed_because_of_TC = ref false in try if not all_instances then match_upats_FO upats env sigma0 ise c; @@ -789,14 +789,14 @@ let rec uniquize = function ws 4 ++ pr_constr_pat env sigma p' ++ fnl () ++ str"of " ++ pr_constr_pat env sigma rule)) : conclude) -type ('ident, 'term) ssrpattern = +type ('ident, 'term) ssrpattern = | T of 'term | In_T of 'term - | X_In_T of 'ident * 'term - | In_X_In_T of 'ident * 'term - | E_In_X_In_T of 'term * 'ident * 'term - | E_As_X_In_T of 'term * 'ident * 'term - + | X_In_T of 'ident * 'term + | In_X_In_T of 'ident * 'term + | E_In_X_In_T of 'term * 'ident * 'term + | E_As_X_In_T of 'term * 'ident * 'term + let pr_pattern = function | T t -> prl_term t | In_T t -> str "in " ++ prl_term t @@ -944,7 +944,7 @@ let of_ftactic ftac gl = in (sigma, ans) -let interp_wit wit ist gl x = +let interp_wit wit ist gl x = let globarg = in_gen (glbwit wit) x in let arg = interp_genarg ist globarg in let (sigma, arg) = of_ftactic arg gl in @@ -1026,9 +1026,9 @@ let interp_pattern ?wit_ssrpatternarg gl red redty = | Evar (k,_) -> if k = h_k || List.mem k acc || Evd.mem sigma0 k then acc else (update k; k::acc) - | _ -> CoqConstr.fold aux acc t in + | _ -> CoqConstr.fold aux acc t in aux [] (nf_evar sigma rp) in - let sigma = + let sigma = List.fold_left (fun sigma e -> if Evd.is_defined sigma e then sigma else (* clear may be recursive *) if Option.is_empty !to_clean then sigma else @@ -1128,7 +1128,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = str " did not instantiate ?" ++ int (Evar.repr e) ++ spc () ++ str "Does the variable bound by the \"in\" construct occur "++ str "in the pattern?") in - let sigma = + let sigma = Evd.add (Evd.remove sigma e) e {e_def with Evd.evar_body = Evar_empty} in sigma, e_body in let ex_value hole = @@ -1138,7 +1138,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = sigma, [pat] in match pattern with | None -> do_subst env0 concl0 concl0 1, UState.empty - | Some (sigma, (T rp | In_T rp)) -> + | Some (sigma, (T rp | In_T rp)) -> let rp = fs sigma rp in let ise = create_evar_defs sigma in let occ = match pattern with Some (_, T _) -> occ | _ -> noindex in @@ -1159,7 +1159,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = let concl = find_T env0 concl0 1 ~k:(fun env c _ h -> let p_sigma = unify_HO env (create_evar_defs sigma) (EConstr.of_constr c) (EConstr.of_constr p) in let sigma, e_body = pop_evar p_sigma ex p in - fs p_sigma (find_X env (fs sigma p) h + fs p_sigma (find_X env (fs sigma p) h ~k:(fun env _ -> do_subst env e_body))) in let _ = end_X () in let _, _, (_, us, _) = end_T () in concl, us @@ -1183,7 +1183,7 @@ let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = | Some (sigma, E_As_X_In_T (e, hole, p)) -> let p, e = fs sigma p, fs sigma e in let ex = ex_value hole in - let rp = + let rp = let e_sigma = unify_HO env0 sigma (EConstr.of_constr hole) (EConstr.of_constr e) in e_sigma, fs e_sigma p in let rp = mk_upat_for ~hack:true env0 sigma0 rp all_ok in @@ -1227,7 +1227,7 @@ let fill_occ_pattern ?raise_NoMatch env sigma cl pat occ h = ;; (* clenup interface for external use *) -let mk_tpattern ?p_origin env sigma0 sigma_t f dir c = +let mk_tpattern ?p_origin env sigma0 sigma_t f dir c = mk_tpattern ?p_origin env sigma0 sigma_t f dir c ;; @@ -1275,7 +1275,7 @@ let is_wildcard ((_, (l, r), _) : cpattern) : bool = match DAst.get l, r with (* "ssrpattern" *) let pr_rpattern = pr_pattern - + let pf_merge_uc uc gl = re_sig (sig_it gl) (Evd.merge_universe_context (project gl) uc) diff --git a/plugins/ssrmatching/ssrmatching.mli b/plugins/ssrmatching/ssrmatching.mli index c6b85738ec..b6ccb4cc6e 100644 --- a/plugins/ssrmatching/ssrmatching.mli +++ b/plugins/ssrmatching/ssrmatching.mli @@ -20,7 +20,7 @@ open Genintern (** Pattern parsing *) -(** The type of context patterns, the patterns of the [set] tactic and +(** The type of context patterns, the patterns of the [set] tactic and [:] tactical. These are patterns that identify a precise subterm. *) type cpattern val pr_cpattern : cpattern -> Pp.t @@ -78,10 +78,10 @@ type occ = (bool * int list) option type subst = env -> constr -> constr -> int -> constr (** [eval_pattern b env sigma t pat occ subst] maps [t] calling [subst] on every - [occ] occurrence of [pat]. The [int] argument is the number of + [occ] occurrence of [pat]. The [int] argument is the number of binders traversed. If [pat] is [None] then then subst is called on [t]. - [t] must live in [env] and [sigma], [pat] must have been interpreted in - (an extension of) [sigma]. + [t] must live in [env] and [sigma], [pat] must have been interpreted in + (an extension of) [sigma]. @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) @return [t] where all [occ] occurrences of [pat] have been mapped using [subst] *) @@ -91,12 +91,12 @@ val eval_pattern : pattern option -> occ -> subst -> constr -(** [fill_occ_pattern b env sigma t pat occ h] is a simplified version of - [eval_pattern]. - It replaces all [occ] occurrences of [pat] in [t] with Rel [h]. - [t] must live in [env] and [sigma], [pat] must have been interpreted in - (an extension of) [sigma]. - @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) +(** [fill_occ_pattern b env sigma t pat occ h] is a simplified version of + [eval_pattern]. + It replaces all [occ] occurrences of [pat] in [t] with Rel [h]. + [t] must live in [env] and [sigma], [pat] must have been interpreted in + (an extension of) [sigma]. + @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) @return the instance of the redex of [pat] that was matched and [t] transformed as described above. *) val fill_occ_pattern : @@ -107,7 +107,7 @@ val fill_occ_pattern : (** *************************** Low level APIs ****************************** *) -(* The primitive matching facility. It matches of a term with holes, like +(* The primitive matching facility. It matches of a term with holes, like the T pattern above, and calls a continuation on its occurrences. *) type ssrdir = L2R | R2L @@ -116,7 +116,7 @@ val pr_dir_side : ssrdir -> Pp.t (** a pattern for a term with wildcards *) type tpattern -(** [mk_tpattern env sigma0 sigma_p ok p_origin dir t] compiles a term [t] +(** [mk_tpattern env sigma0 sigma_p ok p_origin dir t] compiles a term [t] living in [env] [sigma] (an extension of [sigma0]) intro a [tpattern]. The [tpattern] can hold a (proof) term [p] and a diction [dir]. The [ok] callback is used to filter occurrences. @@ -130,14 +130,14 @@ val mk_tpattern : ssrdir -> constr -> evar_map * tpattern -(** [findP env t i k] is a stateful function that finds the next occurrence - of a tpattern and calls the callback [k] to map the subterm matched. - The [int] argument passed to [k] is the number of binders traversed so far - plus the initial value [i]. - @return [t] where the subterms identified by the selected occurrences of +(** [findP env t i k] is a stateful function that finds the next occurrence + of a tpattern and calls the callback [k] to map the subterm matched. + The [int] argument passed to [k] is the number of binders traversed so far + plus the initial value [i]. + @return [t] where the subterms identified by the selected occurrences of the patter have been mapped using [k] @raise NoMatch if the raise_NoMatch flag given to [mk_tpattern_matcher] is - [true] and if the pattern did not match + [true] and if the pattern did not match @raise UserEerror if the raise_NoMatch flag given to [mk_tpattern_matcher] is [false] and if the pattern did not match *) type find_P = @@ -150,11 +150,11 @@ type find_P = type conclude = unit -> constr * ssrdir * (evar_map * UState.t * constr) -(** [mk_tpattern_matcher b o sigma0 occ sigma_tplist] creates a pair +(** [mk_tpattern_matcher b o sigma0 occ sigma_tplist] creates a pair a function [find_P] and [conclude] with the behaviour explained above. The flag [b] (default [false]) changes the error reporting behaviour of [find_P] if none of the [tpattern] matches. The argument [o] can - be passed to tune the [UserError] eventually raised (useful if the + be passed to tune the [UserError] eventually raised (useful if the pattern is coming from the LHS/RHS of an equation) *) val mk_tpattern_matcher : ?all_instances:bool -> @@ -163,24 +163,24 @@ val mk_tpattern_matcher : evar_map -> occ -> evar_map * tpattern list -> find_P * conclude -(** Example of [mk_tpattern_matcher] to implement +(** Example of [mk_tpattern_matcher] to implement [rewrite \{occ\}\[in t\]rules]. - It first matches "in t" (called [pat]), then in all matched subterms + It first matches "in t" (called [pat]), then in all matched subterms it matches the LHS of the rules using [find_R]. [concl0] is the initial goal, [concl] will be the goal where some terms are replaced by a De Bruijn index. The [rw_progress] extra check selects only occurrences that are not rewritten to themselves (e.g. - an occurrence "x + x" rewritten with the commutativity law of addition + an occurrence "x + x" rewritten with the commutativity law of addition is skipped) {[ let find_R, conclude = match pat with | Some (_, In_T _) -> let aux (sigma, pats) (d, r, lhs, rhs) = - let sigma, pat = + let sigma, pat = mk_tpattern env0 sigma0 (sigma, r) (rw_progress rhs) d lhs in sigma, pats @ [pat] in let rpats = List.fold_left aux (r_sigma, []) rules in let find_R, end_R = mk_tpattern_matcher sigma0 occ rpats in - find_R ~k:(fun _ _ h -> mkRel h), + find_R ~k:(fun _ _ h -> mkRel h), fun cl -> let rdx, d, r = end_R () in (d,r),rdx | _ -> ... in let concl = eval_pattern env0 sigma0 concl0 pat occ find_R in @@ -194,7 +194,7 @@ val pf_fill_occ_term : goal sigma -> occ -> evar_map * EConstr.t -> EConstr.t * (* It may be handy to inject a simple term into the first form of cpattern *) val cpattern_of_term : char * glob_constr_and_expr -> Geninterp.interp_sign -> cpattern -(** Helpers to make stateful closures. Example: a [find_P] function may be +(** Helpers to make stateful closures. Example: a [find_P] function may be called many times, but the pattern instantiation phase is performed only the first time. The corresponding [conclude] has to return the instantiated pattern redex. Since it is up to [find_P] to raise [NoMatch] if the pattern |
