From 5b7eafd0f00a16d78f99a27f5c7d5a0de77dc7e6 Mon Sep 17 00:00:00 2001 From: Stephane Glondu Date: Wed, 21 Jul 2010 09:46:51 +0200 Subject: Imported Upstream snapshot 8.3~beta0+13298 --- plugins/funind/Recdef.v | 48 + plugins/funind/functional_principles_proofs.ml | 1710 +++++++++++++++++++++++ plugins/funind/functional_principles_proofs.mli | 19 + plugins/funind/functional_principles_types.ml | 737 ++++++++++ plugins/funind/functional_principles_types.mli | 34 + plugins/funind/g_indfun.ml4 | 524 +++++++ plugins/funind/indfun.ml | 776 ++++++++++ plugins/funind/indfun_common.ml | 558 ++++++++ plugins/funind/indfun_common.mli | 121 ++ plugins/funind/invfun.ml | 1020 ++++++++++++++ plugins/funind/merge.ml | 1032 ++++++++++++++ plugins/funind/rawterm_to_relation.ml | 1419 +++++++++++++++++++ plugins/funind/rawterm_to_relation.mli | 16 + plugins/funind/rawtermops.ml | 718 ++++++++++ plugins/funind/rawtermops.mli | 126 ++ plugins/funind/recdef.ml | 1473 +++++++++++++++++++ plugins/funind/recdef_plugin.mllib | 11 + plugins/funind/vo.itarget | 1 + 18 files changed, 10343 insertions(+) create mode 100644 plugins/funind/Recdef.v create mode 100644 plugins/funind/functional_principles_proofs.ml create mode 100644 plugins/funind/functional_principles_proofs.mli create mode 100644 plugins/funind/functional_principles_types.ml create mode 100644 plugins/funind/functional_principles_types.mli create mode 100644 plugins/funind/g_indfun.ml4 create mode 100644 plugins/funind/indfun.ml create mode 100644 plugins/funind/indfun_common.ml create mode 100644 plugins/funind/indfun_common.mli create mode 100644 plugins/funind/invfun.ml create mode 100644 plugins/funind/merge.ml create mode 100644 plugins/funind/rawterm_to_relation.ml create mode 100644 plugins/funind/rawterm_to_relation.mli create mode 100644 plugins/funind/rawtermops.ml create mode 100644 plugins/funind/rawtermops.mli create mode 100644 plugins/funind/recdef.ml create mode 100644 plugins/funind/recdef_plugin.mllib create mode 100644 plugins/funind/vo.itarget (limited to 'plugins/funind') diff --git a/plugins/funind/Recdef.v b/plugins/funind/Recdef.v new file mode 100644 index 00000000..00302a74 --- /dev/null +++ b/plugins/funind/Recdef.v @@ -0,0 +1,48 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* A) -> A -> A := + fun (fl : A -> A) (def : A) => + match n with + | O => def + | S m => fl (iter m fl def) + end. +End Iter. + +Theorem SSplus_lt : forall p p' : nat, p < S (S (p + p')). + intro p; intro p'; change (S p <= S (S (p + p'))); + apply le_S; apply Gt.gt_le_S; change (p < S (p + p')); + apply Lt.le_lt_n_Sm; apply Plus.le_plus_l. +Qed. + + +Theorem Splus_lt : forall p p' : nat, p' < S (p + p'). + intro p; intro p'; change (S p' <= S (p + p')); + apply Gt.gt_le_S; change (p' < S (p + p')); apply Lt.le_lt_n_Sm; + apply Plus.le_plus_r. +Qed. + +Theorem le_lt_SS : forall x y, x <= y -> x < S (S y). +intro x; intro y; intro H; change (S x <= S (S y)); + apply le_S; apply Gt.gt_le_S; change (x < S y); + apply Lt.le_lt_n_Sm; exact H. +Qed. + +Inductive max_type (m n:nat) : Set := + cmt : forall v, m <= v -> n <= v -> max_type m n. + +Definition max : forall m n:nat, max_type m n. +intros m n; case (Compare_dec.le_gt_dec m n). +intros h; exists n; [exact h | apply le_n]. +intros h; exists m; [apply le_n | apply Lt.lt_le_weak; exact h]. +Defined. diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml new file mode 100644 index 00000000..e2cad944 --- /dev/null +++ b/plugins/funind/functional_principles_proofs.ml @@ -0,0 +1,1710 @@ +open Printer +open Util +open Term +open Termops +open Namegen +open Names +open Declarations +open Pp +open Entries +open Hiddentac +open Evd +open Tacmach +open Proof_type +open Tacticals +open Tactics +open Indfun_common +open Libnames + +let msgnl = Pp.msgnl + + +let observe strm = + if do_observe () + then Pp.msgnl strm + else () + +let observennl strm = + if do_observe () + then begin Pp.msg strm;Pp.pp_flush () end + else () + + + + +let do_observe_tac s tac g = + try let v = tac g in (* msgnl (goal ++ fnl () ++ (str s)++(str " ")++(str "finished")); *) v + with e -> + let goal = begin try (Printer.pr_goal (sig_it g)) with _ -> assert false end in + msgnl (str "observation "++ s++str " raised exception " ++ + Cerrors.explain_exn e ++ str " on goal " ++ goal ); + raise e;; + +let observe_tac_stream s tac g = + if do_observe () + then do_observe_tac s tac g + else tac g + +let observe_tac s tac g = observe_tac_stream (str s) tac g + +(* let tclTRYD tac = *) +(* if !Flags.debug || do_observe () *) +(* then (fun g -> try (\* do_observe_tac "" *\)tac g with _ -> tclIDTAC g) *) +(* else tac *) + + +let list_chop ?(msg="") n l = + try + list_chop n l + with Failure (msg') -> + failwith (msg ^ msg') + + +let make_refl_eq constructor type_of_t t = +(* let refl_equal_term = Lazy.force refl_equal in *) + mkApp(constructor,[|type_of_t;t|]) + + +type pte_info = + { + proving_tac : (identifier list -> Tacmach.tactic); + is_valid : constr -> bool + } + +type ptes_info = pte_info Idmap.t + +type 'a dynamic_info = + { + nb_rec_hyps : int; + rec_hyps : identifier list ; + eq_hyps : identifier list; + info : 'a + } + +type body_info = constr dynamic_info + + +let finish_proof dynamic_infos g = + observe_tac "finish" + ( h_assumption) + g + + +let refine c = + Tacmach.refine_no_check c + +let thin l = + Tacmach.thin_no_check l + + +let cut_replacing id t tac :tactic= + tclTHENS (cut t) + [ tclTHEN (thin_no_check [id]) (introduction_no_check id); + tac + ] + +let intro_erasing id = tclTHEN (thin [id]) (introduction id) + + + +let rec_hyp_id = id_of_string "rec_hyp" + +let is_trivial_eq t = + let res = try + begin + match kind_of_term t with + | App(f,[|_;t1;t2|]) when eq_constr f (Lazy.force eq) -> + eq_constr t1 t2 + | App(f,[|t1;a1;t2;a2|]) when eq_constr f (jmeq ()) -> + eq_constr t1 t2 && eq_constr a1 a2 + | _ -> false + end + with _ -> false + in +(* observe (str "is_trivial_eq " ++ Printer.pr_lconstr t ++ (if res then str " true" else str " false")); *) + res + +let rec incompatible_constructor_terms t1 t2 = + let c1,arg1 = decompose_app t1 + and c2,arg2 = decompose_app t2 + in + (not (eq_constr t1 t2)) && + isConstruct c1 && isConstruct c2 && + ( + not (eq_constr c1 c2) || + List.exists2 incompatible_constructor_terms arg1 arg2 + ) + +let is_incompatible_eq t = + let res = + try + match kind_of_term t with + | App(f,[|_;t1;t2|]) when eq_constr f (Lazy.force eq) -> + incompatible_constructor_terms t1 t2 + | App(f,[|u1;t1;u2;t2|]) when eq_constr f (jmeq ()) -> + (eq_constr u1 u2 && + incompatible_constructor_terms t1 t2) + | _ -> false + with _ -> false + in + if res then observe (str "is_incompatible_eq " ++ Printer.pr_lconstr t); + res + +let change_hyp_with_using msg hyp_id t tac : tactic = + fun g -> + let prov_id = pf_get_new_id hyp_id g in + tclTHENS + ((* observe_tac msg *) (assert_by (Name prov_id) t (tclCOMPLETE tac))) + [tclTHENLIST + [ + (* observe_tac "change_hyp_with_using thin" *) (thin [hyp_id]); + (* observe_tac "change_hyp_with_using rename " *) (h_rename [prov_id,hyp_id]) + ]] g + +exception TOREMOVE + + +let prove_trivial_eq h_id context (constructor,type_of_term,term) = + let nb_intros = List.length context in + tclTHENLIST + [ + tclDO nb_intros intro; (* introducing context *) + (fun g -> + let context_hyps = + fst (list_chop ~msg:"prove_trivial_eq : " nb_intros (pf_ids_of_hyps g)) + in + let context_hyps' = + (mkApp(constructor,[|type_of_term;term|])):: + (List.map mkVar context_hyps) + in + let to_refine = applist(mkVar h_id,List.rev context_hyps') in + refine to_refine g + ) + ] + + + +let find_rectype env c = + let (t, l) = decompose_app (Reduction.whd_betadeltaiota env c) in + match kind_of_term t with + | Ind ind -> (t, l) + | Construct _ -> (t,l) + | _ -> raise Not_found + + +let isAppConstruct ?(env=Global.env ()) t = + try + let t',l = find_rectype (Global.env ()) t in + observe (str "isAppConstruct : " ++ Printer.pr_lconstr t ++ str " -> " ++ Printer.pr_lconstr (applist (t',l))); + true + with Not_found -> false + +let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta *) + let clos_norm_flags flgs env sigma t = + Closure.norm_val (Closure.create_clos_infos flgs env) (Closure.inject (Reductionops.nf_evar sigma t)) in + clos_norm_flags Closure.betaiotazeta Environ.empty_env Evd.empty + + + +let change_eq env sigma hyp_id (context:rel_context) x t end_of_type = + let nochange ?t' msg = + begin + observe (str ("Not treating ( "^msg^" )") ++ pr_lconstr t ++ str " " ++ match t' with None -> str "" | Some t -> Printer.pr_lconstr t ); + failwith "NoChange"; + end + in + let eq_constr = Reductionops.is_conv env sigma in + if not (noccurn 1 end_of_type) + then nochange "dependent"; (* if end_of_type depends on this term we don't touch it *) + if not (isApp t) then nochange "not an equality"; + let f_eq,args = destApp t in + let constructor,t1,t2,t1_typ = + try + if (eq_constr f_eq (Lazy.force eq)) + then + let t1 = (args.(1),args.(0)) + and t2 = (args.(2),args.(0)) + and t1_typ = args.(0) + in + (Lazy.force refl_equal,t1,t2,t1_typ) + else + if (eq_constr f_eq (jmeq ())) + then + (jmeq_refl (),(args.(1),args.(0)),(args.(3),args.(2)),args.(0)) + else nochange "not an equality" + with _ -> nochange "not an equality" + in + if not ((closed0 (fst t1)) && (closed0 (snd t1)))then nochange "not a closed lhs"; + let rec compute_substitution sub t1 t2 = +(* observe (str "compute_substitution : " ++ pr_lconstr t1 ++ str " === " ++ pr_lconstr t2); *) + if isRel t2 + then + let t2 = destRel t2 in + begin + try + let t1' = Intmap.find t2 sub in + if not (eq_constr t1 t1') then nochange "twice bound variable"; + sub + with Not_found -> + assert (closed0 t1); + Intmap.add t2 t1 sub + end + else if isAppConstruct t1 && isAppConstruct t2 + then + begin + let c1,args1 = find_rectype env t1 + and c2,args2 = find_rectype env t2 + in + if not (eq_constr c1 c2) then nochange "cannot solve (diff)"; + List.fold_left2 compute_substitution sub args1 args2 + end + else + if (eq_constr t1 t2) then sub else nochange ~t':(make_refl_eq constructor (Reduction.whd_betadeltaiota env t1) t2) "cannot solve (diff)" + in + let sub = compute_substitution Intmap.empty (snd t1) (snd t2) in + let sub = compute_substitution sub (fst t1) (fst t2) in + let end_of_type_with_pop = pop end_of_type in (*the equation will be removed *) + let new_end_of_type = + (* Ugly hack to prevent Map.fold order change between ocaml-3.08.3 and ocaml-3.08.4 + Can be safely replaced by the next comment for Ocaml >= 3.08.4 + *) + let sub' = Intmap.fold (fun i t acc -> (i,t)::acc) sub [] in + let sub'' = List.sort (fun (x,_) (y,_) -> Pervasives.compare x y) sub' in + List.fold_left (fun end_of_type (i,t) -> lift 1 (substnl [t] (i-1) end_of_type)) + end_of_type_with_pop + sub'' + in + let old_context_length = List.length context + 1 in + let witness_fun = + mkLetIn(Anonymous,make_refl_eq constructor t1_typ (fst t1),t, + mkApp(mkVar hyp_id,Array.init old_context_length (fun i -> mkRel (old_context_length - i))) + ) + in + let new_type_of_hyp,ctxt_size,witness_fun = + list_fold_left_i + (fun i (end_of_type,ctxt_size,witness_fun) ((x',b',t') as decl) -> + try + let witness = Intmap.find i sub in + if b' <> None then anomaly "can not redefine a rel!"; + (pop end_of_type,ctxt_size,mkLetIn(x',witness,t',witness_fun)) + with Not_found -> + (mkProd_or_LetIn decl end_of_type, ctxt_size + 1, mkLambda_or_LetIn decl witness_fun) + ) + 1 + (new_end_of_type,0,witness_fun) + context + in + let new_type_of_hyp = + Reductionops.nf_betaiota Evd.empty new_type_of_hyp in + let new_ctxt,new_end_of_type = + decompose_prod_n_assum ctxt_size new_type_of_hyp + in + let prove_new_hyp : tactic = + tclTHEN + (tclDO ctxt_size intro) + (fun g -> + let all_ids = pf_ids_of_hyps g in + let new_ids,_ = list_chop ctxt_size all_ids in + let to_refine = applist(witness_fun,List.rev_map mkVar new_ids) in + refine to_refine g + ) + in + let simpl_eq_tac = + change_hyp_with_using "prove_pattern_simplification" hyp_id new_type_of_hyp prove_new_hyp + in +(* observe (str "In " ++ Ppconstr.pr_id hyp_id ++ *) +(* str "removing an equation " ++ fnl ()++ *) +(* str "old_typ_of_hyp :=" ++ *) +(* Printer.pr_lconstr_env *) +(* env *) +(* (it_mkProd_or_LetIn ~init:end_of_type ((x,None,t)::context)) *) +(* ++ fnl () ++ *) +(* str "new_typ_of_hyp := "++ *) +(* Printer.pr_lconstr_env env new_type_of_hyp ++ fnl () *) +(* ++ str "old context := " ++ pr_rel_context env context ++ fnl () *) +(* ++ str "new context := " ++ pr_rel_context env new_ctxt ++ fnl () *) +(* ++ str "old type := " ++ pr_lconstr end_of_type ++ fnl () *) +(* ++ str "new type := " ++ pr_lconstr new_end_of_type ++ fnl () *) +(* ); *) + new_ctxt,new_end_of_type,simpl_eq_tac + + +let is_property ptes_info t_x full_type_of_hyp = + if isApp t_x + then + let pte,args = destApp t_x in + if isVar pte && array_for_all closed0 args + then + try + let info = Idmap.find (destVar pte) ptes_info in + info.is_valid full_type_of_hyp + with Not_found -> false + else false + else false + +let isLetIn t = + match kind_of_term t with + | LetIn _ -> true + | _ -> false + + +let h_reduce_with_zeta = + h_reduce + (Rawterm.Cbv + {Rawterm.all_flags + with Rawterm.rDelta = false; + }) + + + +let rewrite_until_var arg_num eq_ids : tactic = + (* tests if the declares recursive argument is neither a Constructor nor + an applied Constructor since such a form for the recursive argument + will break the Guard when trying to save the Lemma. + *) + let test_var g = + let _,args = destApp (pf_concl g) in + not ((isConstruct args.(arg_num)) || isAppConstruct args.(arg_num)) + in + let rec do_rewrite eq_ids g = + if test_var g + then tclIDTAC g + else + match eq_ids with + | [] -> anomaly "Cannot find a way to prove recursive property"; + | eq_id::eq_ids -> + tclTHEN + (tclTRY (Equality.rewriteRL (mkVar eq_id))) + (do_rewrite eq_ids) + g + in + do_rewrite eq_ids + + +let rec_pte_id = id_of_string "Hrec" +let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma = + let coq_False = Coqlib.build_coq_False () in + let coq_True = Coqlib.build_coq_True () in + let coq_I = Coqlib.build_coq_I () in + let rec scan_type context type_of_hyp : tactic = + if isLetIn type_of_hyp then + let real_type_of_hyp = it_mkProd_or_LetIn ~init:type_of_hyp context in + let reduced_type_of_hyp = nf_betaiotazeta real_type_of_hyp in + (* length of context didn't change ? *) + let new_context,new_typ_of_hyp = + decompose_prod_n_assum (List.length context) reduced_type_of_hyp + in + tclTHENLIST + [ + h_reduce_with_zeta + (Tacticals.onHyp hyp_id) + ; + scan_type new_context new_typ_of_hyp + + ] + else if isProd type_of_hyp + then + begin + let (x,t_x,t') = destProd type_of_hyp in + let actual_real_type_of_hyp = it_mkProd_or_LetIn ~init:t' context in + if is_property ptes_infos t_x actual_real_type_of_hyp then + begin + let pte,pte_args = (destApp t_x) in + let (* fix_info *) prove_rec_hyp = (Idmap.find (destVar pte) ptes_infos).proving_tac in + let popped_t' = pop t' in + let real_type_of_hyp = it_mkProd_or_LetIn ~init:popped_t' context in + let prove_new_type_of_hyp = + let context_length = List.length context in + tclTHENLIST + [ + tclDO context_length intro; + (fun g -> + let context_hyps_ids = + fst (list_chop ~msg:"rec hyp : context_hyps" + context_length (pf_ids_of_hyps g)) + in + let rec_pte_id = pf_get_new_id rec_pte_id g in + let to_refine = + applist(mkVar hyp_id, + List.rev_map mkVar (rec_pte_id::context_hyps_ids) + ) + in +(* observe_tac "rec hyp " *) + (tclTHENS + (assert_tac (Name rec_pte_id) t_x) + [ + (* observe_tac "prove rec hyp" *) (prove_rec_hyp eq_hyps); +(* observe_tac "prove rec hyp" *) + (refine to_refine) + ]) + g + ) + ] + in + tclTHENLIST + [ +(* observe_tac "hyp rec" *) + (change_hyp_with_using "rec_hyp_tac" hyp_id real_type_of_hyp prove_new_type_of_hyp); + scan_type context popped_t' + ] + end + else if eq_constr t_x coq_False then + begin +(* observe (str "Removing : "++ Ppconstr.pr_id hyp_id++ *) +(* str " since it has False in its preconds " *) +(* ); *) + raise TOREMOVE; (* False -> .. useless *) + end + else if is_incompatible_eq t_x then raise TOREMOVE (* t_x := C1 ... = C2 ... *) + else if eq_constr t_x coq_True (* Trivial => we remove this precons *) + then +(* observe (str "In "++Ppconstr.pr_id hyp_id++ *) +(* str " removing useless precond True" *) +(* ); *) + let popped_t' = pop t' in + let real_type_of_hyp = + it_mkProd_or_LetIn ~init:popped_t' context + in + let prove_trivial = + let nb_intro = List.length context in + tclTHENLIST [ + tclDO nb_intro intro; + (fun g -> + let context_hyps = + fst (list_chop ~msg:"removing True : context_hyps "nb_intro (pf_ids_of_hyps g)) + in + let to_refine = + applist (mkVar hyp_id, + List.rev (coq_I::List.map mkVar context_hyps) + ) + in + refine to_refine g + ) + ] + in + tclTHENLIST[ + change_hyp_with_using "prove_trivial" hyp_id real_type_of_hyp + ((* observe_tac "prove_trivial" *) prove_trivial); + scan_type context popped_t' + ] + else if is_trivial_eq t_x + then (* t_x := t = t => we remove this precond *) + let popped_t' = pop t' in + let real_type_of_hyp = + it_mkProd_or_LetIn ~init:popped_t' context + in + let hd,args = destApp t_x in + let get_args hd args = + if eq_constr hd (Lazy.force eq) + then (Lazy.force refl_equal,args.(0),args.(1)) + else (jmeq_refl (),args.(0),args.(1)) + in + tclTHENLIST + [ + change_hyp_with_using + "prove_trivial_eq" + hyp_id + real_type_of_hyp + ((* observe_tac "prove_trivial_eq" *) + (prove_trivial_eq hyp_id context (get_args hd args))); + scan_type context popped_t' + ] + else + begin + try + let new_context,new_t',tac = change_eq env sigma hyp_id context x t_x t' in + tclTHEN + tac + (scan_type new_context new_t') + with Failure "NoChange" -> + (* Last thing todo : push the rel in the context and continue *) + scan_type ((x,None,t_x)::context) t' + end + end + else + tclIDTAC + in + try + scan_type [] (Typing.type_of env sigma (mkVar hyp_id)), [hyp_id] + with TOREMOVE -> + thin [hyp_id],[] + + +let clean_goal_with_heq ptes_infos continue_tac dyn_infos = + fun g -> + let env = pf_env g + and sigma = project g + in + let tac,new_hyps = + List.fold_left ( + fun (hyps_tac,new_hyps) hyp_id -> + let hyp_tac,new_hyp = + clean_hyp_with_heq ptes_infos dyn_infos.eq_hyps hyp_id env sigma + in + (tclTHEN hyp_tac hyps_tac),new_hyp@new_hyps + ) + (tclIDTAC,[]) + dyn_infos.rec_hyps + in + let new_infos = + { dyn_infos with + rec_hyps = new_hyps; + nb_rec_hyps = List.length new_hyps + } + in + tclTHENLIST + [ + tac ; + (* observe_tac "clean_hyp_with_heq continue" *) (continue_tac new_infos) + ] + g + +let heq_id = id_of_string "Heq" + +let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos = + fun g -> + let nb_first_intro = nb_prod - 1 - dyn_infos.nb_rec_hyps in + tclTHENLIST + [ + (* We first introduce the variables *) + tclDO nb_first_intro (intro_avoiding dyn_infos.rec_hyps); + (* Then the equation itself *) + intro_using heq_id; + onLastHypId (fun heq_id -> tclTHENLIST [ + (* Then the new hypothesis *) + tclMAP introduction_no_check dyn_infos.rec_hyps; + (* observe_tac "after_introduction" *)(fun g' -> + (* We get infos on the equations introduced*) + let new_term_value_eq = pf_type_of g' (mkVar heq_id) in + (* compute the new value of the body *) + let new_term_value = + match kind_of_term new_term_value_eq with + | App(f,[| _;_;args2 |]) -> args2 + | _ -> + observe (str "cannot compute new term value : " ++ pr_gls g' ++ fnl () ++ str "last hyp is" ++ + pr_lconstr_env (pf_env g') new_term_value_eq + ); + anomaly "cannot compute new term value" + in + let fun_body = + mkLambda(Anonymous, + pf_type_of g' term, + replace_term term (mkRel 1) dyn_infos.info + ) + in + let new_body = pf_nf_betaiota g' (mkApp(fun_body,[| new_term_value |])) in + let new_infos = + {dyn_infos with + info = new_body; + eq_hyps = heq_id::dyn_infos.eq_hyps + } + in + clean_goal_with_heq ptes_infos continue_tac new_infos g' + )]) + ] + g + + +let my_orelse tac1 tac2 g = + try + tac1 g + with e -> +(* observe (str "using snd tac since : " ++ Cerrors.explain_exn e); *) + tac2 g + +let instanciate_hyps_with_args (do_prove:identifier list -> tactic) hyps args_id = + let args = Array.of_list (List.map mkVar args_id) in + let instanciate_one_hyp hid = + my_orelse + ( (* we instanciate the hyp if possible *) + fun g -> + let prov_hid = pf_get_new_id hid g in + tclTHENLIST[ + pose_proof (Name prov_hid) (mkApp(mkVar hid,args)); + thin [hid]; + h_rename [prov_hid,hid] + ] g + ) + ( (* + if not then we are in a mutual function block + and this hyp is a recursive hyp on an other function. + + We are not supposed to use it while proving this + principle so that we can trash it + + *) + (fun g -> +(* observe (str "Instanciation: removing hyp " ++ Ppconstr.pr_id hid); *) + thin [hid] g + ) + ) + in + if args_id = [] + then + tclTHENLIST [ + tclMAP (fun hyp_id -> h_reduce_with_zeta (Tacticals.onHyp hyp_id)) hyps; + do_prove hyps + ] + else + tclTHENLIST + [ + tclMAP (fun hyp_id -> h_reduce_with_zeta (Tacticals.onHyp hyp_id)) hyps; + tclMAP instanciate_one_hyp hyps; + (fun g -> + let all_g_hyps_id = + List.fold_right Idset.add (pf_ids_of_hyps g) Idset.empty + in + let remaining_hyps = + List.filter (fun id -> Idset.mem id all_g_hyps_id) hyps + in + do_prove remaining_hyps g + ) + ] + +let build_proof + (interactive_proof:bool) + (fnames:constant list) + ptes_infos + dyn_infos + : tactic = + let rec build_proof_aux do_finalize dyn_infos : tactic = + fun g -> +(* observe (str "proving on " ++ Printer.pr_lconstr_env (pf_env g) term);*) + match kind_of_term dyn_infos.info with + | Case(ci,ct,t,cb) -> + let do_finalize_t dyn_info' = + fun g -> + let t = dyn_info'.info in + let dyn_infos = {dyn_info' with info = + mkCase(ci,ct,t,cb)} in + let g_nb_prod = nb_prod (pf_concl g) in + let type_of_term = pf_type_of g t in + let term_eq = + make_refl_eq (Lazy.force refl_equal) type_of_term t + in + tclTHENSEQ + [ + h_generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps)); + thin dyn_infos.rec_hyps; + pattern_option [(false,[1]),t] None; + (fun g -> observe_tac "toto" ( + tclTHENSEQ [h_simplest_case t; + (fun g' -> + let g'_nb_prod = nb_prod (pf_concl g') in + let nb_instanciate_partial = g'_nb_prod - g_nb_prod in + observe_tac "treat_new_case" + (treat_new_case + ptes_infos + nb_instanciate_partial + (build_proof do_finalize) + t + dyn_infos) + g' + ) + + ]) g + ) + ] + g + in + build_proof do_finalize_t {dyn_infos with info = t} g + | Lambda(n,t,b) -> + begin + match kind_of_term( pf_concl g) with + | Prod _ -> + tclTHEN + intro + (fun g' -> + let (id,_,_) = pf_last_hyp g' in + let new_term = + pf_nf_betaiota g' + (mkApp(dyn_infos.info,[|mkVar id|])) + in + let new_infos = {dyn_infos with info = new_term} in + let do_prove new_hyps = + build_proof do_finalize + {new_infos with + rec_hyps = new_hyps; + nb_rec_hyps = List.length new_hyps + } + in +(* observe_tac "Lambda" *) (instanciate_hyps_with_args do_prove new_infos.rec_hyps [id]) g' + (* build_proof do_finalize new_infos g' *) + ) g + | _ -> + do_finalize dyn_infos g + end + | Cast(t,_,_) -> + build_proof do_finalize {dyn_infos with info = t} g + | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ -> + do_finalize dyn_infos g + | App(_,_) -> + let f,args = decompose_app dyn_infos.info in + begin + match kind_of_term f with + | App _ -> assert false (* we have collected all the app in decompose_app *) + | Var _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _ | Sort _ | Prod _ -> + let new_infos = + { dyn_infos with + info = (f,args) + } + in + build_proof_args do_finalize new_infos g + | Const c when not (List.mem c fnames) -> + let new_infos = + { dyn_infos with + info = (f,args) + } + in +(* Pp.msgnl (str "proving in " ++ pr_lconstr_env (pf_env g) dyn_infos.info); *) + build_proof_args do_finalize new_infos g + | Const _ -> + do_finalize dyn_infos g + | Lambda _ -> + let new_term = + Reductionops.nf_beta Evd.empty dyn_infos.info in + build_proof do_finalize {dyn_infos with info = new_term} + g + | LetIn _ -> + let new_infos = + { dyn_infos with info = nf_betaiotazeta dyn_infos.info } + in + + tclTHENLIST + [tclMAP + (fun hyp_id -> h_reduce_with_zeta (Tacticals.onHyp hyp_id)) + dyn_infos.rec_hyps; + h_reduce_with_zeta Tacticals.onConcl; + build_proof do_finalize new_infos + ] + g + | Cast(b,_,_) -> + build_proof do_finalize {dyn_infos with info = b } g + | Case _ | Fix _ | CoFix _ -> + let new_finalize dyn_infos = + let new_infos = + { dyn_infos with + info = dyn_infos.info,args + } + in + build_proof_args do_finalize new_infos + in + build_proof new_finalize {dyn_infos with info = f } g + end + | Fix _ | CoFix _ -> + error ( "Anonymous local (co)fixpoints are not handled yet") + + | Prod _ -> error "Prod" + | LetIn _ -> + let new_infos = + { dyn_infos with + info = nf_betaiotazeta dyn_infos.info + } + in + + tclTHENLIST + [tclMAP + (fun hyp_id -> h_reduce_with_zeta (Tacticals.onHyp hyp_id)) + dyn_infos.rec_hyps; + h_reduce_with_zeta Tacticals.onConcl; + build_proof do_finalize new_infos + ] g + | Rel _ -> anomaly "Free var in goal conclusion !" + and build_proof do_finalize dyn_infos g = +(* observe (str "proving with "++Printer.pr_lconstr dyn_infos.info++ str " on goal " ++ pr_gls g); *) + observe_tac "build_proof" (build_proof_aux do_finalize dyn_infos) g + and build_proof_args do_finalize dyn_infos (* f_args' args *) :tactic = + fun g -> + let (f_args',args) = dyn_infos.info in + let tac : tactic = + fun g -> + match args with + | [] -> + do_finalize {dyn_infos with info = f_args'} g + | arg::args -> +(* observe (str "build_proof_args with arg := "++ pr_lconstr_env (pf_env g) arg++ *) +(* fnl () ++ *) +(* pr_goal (Tacmach.sig_it g) *) +(* ); *) + let do_finalize dyn_infos = + let new_arg = dyn_infos.info in + (* tclTRYD *) + (build_proof_args + do_finalize + {dyn_infos with info = (mkApp(f_args',[|new_arg|])), args} + ) + in + build_proof do_finalize + {dyn_infos with info = arg } + g + in + (* observe_tac "build_proof_args" *) (tac ) g + in + let do_finish_proof dyn_infos = + (* tclTRYD *) (clean_goal_with_heq + ptes_infos + finish_proof dyn_infos) + in + (* observe_tac "build_proof" *) + (build_proof (clean_goal_with_heq ptes_infos do_finish_proof) dyn_infos) + + + + + + + + + + + + +(* Proof of principles from structural functions *) +let is_pte_type t = + isSort ((strip_prod t)) + +let is_pte (_,_,t) = is_pte_type t + + + + +type static_fix_info = + { + idx : int; + name : identifier; + types : types; + offset : int; + nb_realargs : int; + body_with_param : constr; + num_in_block : int + } + + + +let prove_rec_hyp_for_struct fix_info = + (fun eq_hyps -> tclTHEN + (rewrite_until_var (fix_info.idx) eq_hyps) + (fun g -> + let _,pte_args = destApp (pf_concl g) in + let rec_hyp_proof = + mkApp(mkVar fix_info.name,array_get_start pte_args) + in + refine rec_hyp_proof g + )) + +let prove_rec_hyp fix_info = + { proving_tac = prove_rec_hyp_for_struct fix_info + ; + is_valid = fun _ -> true + } + + +exception Not_Rec + +let generalize_non_dep hyp g = +(* observe (str "rec id := " ++ Ppconstr.pr_id hyp); *) + let hyps = [hyp] in + let env = Global.env () in + let hyp_typ = pf_type_of g (mkVar hyp) in + let to_revert,_ = + Environ.fold_named_context_reverse (fun (clear,keep) (hyp,_,_ as decl) -> + if List.mem hyp hyps + or List.exists (occur_var_in_decl env hyp) keep + or occur_var env hyp hyp_typ + or Termops.is_section_variable hyp (* should be dangerous *) + then (clear,decl::keep) + else (hyp::clear,keep)) + ~init:([],[]) (pf_env g) + in +(* observe (str "to_revert := " ++ prlist_with_sep spc Ppconstr.pr_id to_revert); *) + tclTHEN + ((* observe_tac "h_generalize" *) (h_generalize (List.map mkVar to_revert) )) + ((* observe_tac "thin" *) (thin to_revert)) + g + +let id_of_decl (na,_,_) = (Nameops.out_name na) +let var_of_decl decl = mkVar (id_of_decl decl) +let revert idl = + tclTHEN + (generalize (List.map mkVar idl)) + (thin idl) + +let generate_equation_lemma fnames f fun_num nb_params nb_args rec_args_num = +(* observe (str "nb_args := " ++ str (string_of_int nb_args)); *) +(* observe (str "nb_params := " ++ str (string_of_int nb_params)); *) +(* observe (str "rec_args_num := " ++ str (string_of_int (rec_args_num + 1) )); *) + let f_def = Global.lookup_constant (destConst f) in + let eq_lhs = mkApp(f,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i))) in + let f_body = + force (Option.get f_def.const_body) + in + let params,f_body_with_params = decompose_lam_n nb_params f_body in + let (_,num),(_,_,bodies) = destFix f_body_with_params in + let fnames_with_params = + let params = Array.init nb_params (fun i -> mkRel(nb_params - i)) in + let fnames = List.rev (Array.to_list (Array.map (fun f -> mkApp(f,params)) fnames)) in + fnames + in +(* observe (str "fnames_with_params " ++ prlist_with_sep fnl pr_lconstr fnames_with_params); *) +(* observe (str "body " ++ pr_lconstr bodies.(num)); *) + let f_body_with_params_and_other_fun = substl fnames_with_params bodies.(num) in +(* observe (str "f_body_with_params_and_other_fun " ++ pr_lconstr f_body_with_params_and_other_fun); *) + let eq_rhs = nf_betaiotazeta (mkApp(compose_lam params f_body_with_params_and_other_fun,Array.init (nb_params + nb_args) (fun i -> mkRel(nb_params + nb_args - i)))) in +(* observe (str "eq_rhs " ++ pr_lconstr eq_rhs); *) + let type_ctxt,type_of_f = decompose_prod_n_assum (nb_params + nb_args) + (Typeops.type_of_constant_type (Global.env()) f_def.const_type) in + let eqn = mkApp(Lazy.force eq,[|type_of_f;eq_lhs;eq_rhs|]) in + let lemma_type = it_mkProd_or_LetIn ~init:eqn type_ctxt in + let f_id = id_of_label (con_label (destConst f)) in + let prove_replacement = + tclTHENSEQ + [ + tclDO (nb_params + rec_args_num + 1) intro; + (* observe_tac "" *) (fun g -> + let rec_id = pf_nth_hyp_id g 1 in + tclTHENSEQ + [(* observe_tac "generalize_non_dep in generate_equation_lemma" *) (generalize_non_dep rec_id); + (* observe_tac "h_case" *) (h_case false (mkVar rec_id,Rawterm.NoBindings)); + intros_reflexivity] g + ) + ] + in + Lemmas.start_proof + (*i The next call to mk_equation_id is valid since we are constructing the lemma + Ensures by: obvious + i*) + (mk_equation_id f_id) + (Decl_kinds.Global,(Decl_kinds.Proof Decl_kinds.Theorem)) + lemma_type + (fun _ _ -> ()); + Pfedit.by (prove_replacement); + Lemmas.save_named false + + + + +let do_replace params rec_arg_num rev_args_id f fun_num all_funs g = + let equation_lemma = + try + let finfos = find_Function_infos (destConst f) in + mkConst (Option.get finfos.equation_lemma) + with (Not_found | Option.IsNone as e) -> + let f_id = id_of_label (con_label (destConst f)) in + (*i The next call to mk_equation_id is valid since we will construct the lemma + Ensures by: obvious + i*) + let equation_lemma_id = (mk_equation_id f_id) in + generate_equation_lemma all_funs f fun_num (List.length params) (List.length rev_args_id) rec_arg_num; + let _ = + match e with + | Option.IsNone -> + let finfos = find_Function_infos (destConst f) in + update_Function + {finfos with + equation_lemma = Some (match Nametab.locate (qualid_of_ident equation_lemma_id) with + ConstRef c -> c + | _ -> Util.anomaly "Not a constant" + ) + } + | _ -> () + + in + Tacinterp.constr_of_id (pf_env g) equation_lemma_id + in + let nb_intro_to_do = nb_prod (pf_concl g) in + tclTHEN + (tclDO nb_intro_to_do intro) + ( + fun g' -> + let just_introduced = nLastDecls nb_intro_to_do g' in + let just_introduced_id = List.map (fun (id,_,_) -> id) just_introduced in + tclTHEN (Equality.rewriteLR equation_lemma) (revert just_introduced_id) g' + ) + g + +let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams : tactic = + fun g -> + let princ_type = pf_concl g in + let princ_info = compute_elim_sig princ_type in + let fresh_id = + let avoid = ref (pf_ids_of_hyps g) in + (fun na -> + let new_id = + match na with + Name id -> fresh_id !avoid (string_of_id id) + | Anonymous -> fresh_id !avoid "H" + in + avoid := new_id :: !avoid; + (Name new_id) + ) + in + let fresh_decl = + (fun (na,b,t) -> + (fresh_id na,b,t) + ) + in + let princ_info : elim_scheme = + { princ_info with + params = List.map fresh_decl princ_info.params; + predicates = List.map fresh_decl princ_info.predicates; + branches = List.map fresh_decl princ_info.branches; + args = List.map fresh_decl princ_info.args + } + in + let get_body const = + match (Global.lookup_constant const ).const_body with + | Some b -> + let body = force b in + Tacred.cbv_norm_flags + (Closure.RedFlags.mkflags [Closure.RedFlags.fZETA]) + (Global.env ()) + (Evd.empty) + body + | None -> error ( "Cannot define a principle over an axiom ") + in + let fbody = get_body fnames.(fun_num) in + let f_ctxt,f_body = decompose_lam fbody in + let f_ctxt_length = List.length f_ctxt in + let diff_params = princ_info.nparams - f_ctxt_length in + let full_params,princ_params,fbody_with_full_params = + if diff_params > 0 + then + let princ_params,full_params = + list_chop diff_params princ_info.params + in + (full_params, (* real params *) + princ_params, (* the params of the principle which are not params of the function *) + substl (* function instanciated with real params *) + (List.map var_of_decl full_params) + f_body + ) + else + let f_ctxt_other,f_ctxt_params = + list_chop (- diff_params) f_ctxt in + let f_body = compose_lam f_ctxt_other f_body in + (princ_info.params, (* real params *) + [],(* all params are full params *) + substl (* function instanciated with real params *) + (List.map var_of_decl princ_info.params) + f_body + ) + in +(* observe (str "full_params := " ++ *) +(* prlist_with_sep spc (fun (na,_,_) -> Ppconstr.pr_id (Nameops.out_name na)) *) +(* full_params *) +(* ); *) +(* observe (str "princ_params := " ++ *) +(* prlist_with_sep spc (fun (na,_,_) -> Ppconstr.pr_id (Nameops.out_name na)) *) +(* princ_params *) +(* ); *) +(* observe (str "fbody_with_full_params := " ++ *) +(* pr_lconstr fbody_with_full_params *) +(* ); *) + let all_funs_with_full_params = + Array.map (fun f -> applist(f, List.rev_map var_of_decl full_params)) all_funs + in + let fix_offset = List.length princ_params in + let ptes_to_fix,infos = + match kind_of_term fbody_with_full_params with + | Fix((idxs,i),(names,typess,bodies)) -> + let bodies_with_all_params = + Array.map + (fun body -> + Reductionops.nf_betaiota Evd.empty + (applist(substl (List.rev (Array.to_list all_funs_with_full_params)) body, + List.rev_map var_of_decl princ_params)) + ) + bodies + in + let info_array = + Array.mapi + (fun i types -> + let types = prod_applist types (List.rev_map var_of_decl princ_params) in + { idx = idxs.(i) - fix_offset; + name = Nameops.out_name (fresh_id names.(i)); + types = types; + offset = fix_offset; + nb_realargs = + List.length + (fst (decompose_lam bodies.(i))) - fix_offset; + body_with_param = bodies_with_all_params.(i); + num_in_block = i + } + ) + typess + in + let pte_to_fix,rev_info = + list_fold_left_i + (fun i (acc_map,acc_info) (pte,_,_) -> + let infos = info_array.(i) in + let type_args,_ = decompose_prod infos.types in + let nargs = List.length type_args in + let f = applist(mkConst fnames.(i), List.rev_map var_of_decl princ_info.params) in + let first_args = Array.init nargs (fun i -> mkRel (nargs -i)) in + let app_f = mkApp(f,first_args) in + let pte_args = (Array.to_list first_args)@[app_f] in + let app_pte = applist(mkVar (Nameops.out_name pte),pte_args) in + let body_with_param,num = + let body = get_body fnames.(i) in + let body_with_full_params = + Reductionops.nf_betaiota Evd.empty ( + applist(body,List.rev_map var_of_decl full_params)) + in + match kind_of_term body_with_full_params with + | Fix((_,num),(_,_,bs)) -> + Reductionops.nf_betaiota Evd.empty + ( + (applist + (substl + (List.rev + (Array.to_list all_funs_with_full_params)) + bs.(num), + List.rev_map var_of_decl princ_params)) + ),num + | _ -> error "Not a mutual block" + in + let info = + {infos with + types = compose_prod type_args app_pte; + body_with_param = body_with_param; + num_in_block = num + } + in +(* observe (str "binding " ++ Ppconstr.pr_id (Nameops.out_name pte) ++ *) +(* str " to " ++ Ppconstr.pr_id info.name); *) + (Idmap.add (Nameops.out_name pte) info acc_map,info::acc_info) + ) + 0 + (Idmap.empty,[]) + (List.rev princ_info.predicates) + in + pte_to_fix,List.rev rev_info + | _ -> Idmap.empty,[] + in + let mk_fixes : tactic = + let pre_info,infos = list_chop fun_num infos in + match pre_info,infos with + | [],[] -> tclIDTAC + | _, this_fix_info::others_infos -> + let other_fix_infos = + List.map + (fun fi -> fi.name,fi.idx + 1 ,fi.types) + (pre_info@others_infos) + in + if other_fix_infos = [] + then + (* observe_tac ("h_fix") *) (h_fix (Some this_fix_info.name) (this_fix_info.idx +1)) + else + h_mutual_fix false this_fix_info.name (this_fix_info.idx + 1) + other_fix_infos + | _ -> anomaly "Not a valid information" + in + let first_tac : tactic = (* every operations until fix creations *) + tclTHENSEQ + [ (* observe_tac "introducing params" *) (intros_using (List.rev_map id_of_decl princ_info.params)); + (* observe_tac "introducing predictes" *) (intros_using (List.rev_map id_of_decl princ_info.predicates)); + (* observe_tac "introducing branches" *) (intros_using (List.rev_map id_of_decl princ_info.branches)); + (* observe_tac "building fixes" *) mk_fixes; + ] + in + let intros_after_fixes : tactic = + fun gl -> + let ctxt,pte_app = (decompose_prod_assum (pf_concl gl)) in + let pte,pte_args = (decompose_app pte_app) in + try + let pte = try destVar pte with _ -> anomaly "Property is not a variable" in + let fix_info = Idmap.find pte ptes_to_fix in + let nb_args = fix_info.nb_realargs in + tclTHENSEQ + [ + (* observe_tac ("introducing args") *) (tclDO nb_args intro); + (fun g -> (* replacement of the function by its body *) + let args = nLastDecls nb_args g in + let fix_body = fix_info.body_with_param in +(* observe (str "fix_body := "++ pr_lconstr_env (pf_env gl) fix_body); *) + let args_id = List.map (fun (id,_,_) -> id) args in + let dyn_infos = + { + nb_rec_hyps = -100; + rec_hyps = []; + info = + Reductionops.nf_betaiota Evd.empty + (applist(fix_body,List.rev_map mkVar args_id)); + eq_hyps = [] + } + in + tclTHENSEQ + [ +(* observe_tac "do_replace" *) + (do_replace + full_params + (fix_info.idx + List.length princ_params) + (args_id@(List.map (fun (id,_,_) -> Nameops.out_name id ) princ_params)) + (all_funs.(fix_info.num_in_block)) + fix_info.num_in_block + all_funs + ); +(* observe_tac "do_replace" *) +(* (do_replace princ_info.params fix_info.idx args_id *) +(* (List.hd (List.rev pte_args)) fix_body); *) + let do_prove = + build_proof + interactive_proof + (Array.to_list fnames) + (Idmap.map prove_rec_hyp ptes_to_fix) + in + let prove_tac branches = + let dyn_infos = + {dyn_infos with + rec_hyps = branches; + nb_rec_hyps = List.length branches + } + in + observe_tac "cleaning" (clean_goal_with_heq + (Idmap.map prove_rec_hyp ptes_to_fix) + do_prove + dyn_infos) + in +(* observe (str "branches := " ++ *) +(* prlist_with_sep spc (fun decl -> Ppconstr.pr_id (id_of_decl decl)) princ_info.branches ++ fnl () ++ *) +(* str "args := " ++ prlist_with_sep spc Ppconstr.pr_id args_id *) + +(* ); *) + (* observe_tac "instancing" *) (instanciate_hyps_with_args prove_tac + (List.rev_map id_of_decl princ_info.branches) + (List.rev args_id)) + ] + g + ); + ] gl + with Not_found -> + let nb_args = min (princ_info.nargs) (List.length ctxt) in + tclTHENSEQ + [ + tclDO nb_args intro; + (fun g -> (* replacement of the function by its body *) + let args = nLastDecls nb_args g in + let args_id = List.map (fun (id,_,_) -> id) args in + let dyn_infos = + { + nb_rec_hyps = -100; + rec_hyps = []; + info = + Reductionops.nf_betaiota Evd.empty + (applist(fbody_with_full_params, + (List.rev_map var_of_decl princ_params)@ + (List.rev_map mkVar args_id) + )); + eq_hyps = [] + } + in + let fname = destConst (fst (decompose_app (List.hd (List.rev pte_args)))) in + tclTHENSEQ + [unfold_in_concl [(all_occurrences,Names.EvalConstRef fname)]; + let do_prove = + build_proof + interactive_proof + (Array.to_list fnames) + (Idmap.map prove_rec_hyp ptes_to_fix) + in + let prove_tac branches = + let dyn_infos = + {dyn_infos with + rec_hyps = branches; + nb_rec_hyps = List.length branches + } + in + clean_goal_with_heq + (Idmap.map prove_rec_hyp ptes_to_fix) + do_prove + dyn_infos + in + instanciate_hyps_with_args prove_tac + (List.rev_map id_of_decl princ_info.branches) + (List.rev args_id) + ] + g + ) + ] + gl + in + tclTHEN + first_tac + intros_after_fixes + g + + + + + + +(* Proof of principles of general functions *) +let h_id = Recdef.h_id +and hrec_id = Recdef.hrec_id +and acc_inv_id = Recdef.acc_inv_id +and ltof_ref = Recdef.ltof_ref +and acc_rel = Recdef.acc_rel +and well_founded = Recdef.well_founded +and delayed_force = Recdef.delayed_force +and h_intros = Recdef.h_intros +and list_rewrite = Recdef.list_rewrite +and evaluable_of_global_reference = Recdef.evaluable_of_global_reference + + + + + +let prove_with_tcc tcc_lemma_constr eqs : tactic = + match !tcc_lemma_constr with + | None -> anomaly "No tcc proof !!" + | Some lemma -> + fun gls -> +(* let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in *) +(* let ids = hid::pf_ids_of_hyps gls in *) + tclTHENSEQ + [ +(* generalize [lemma]; *) +(* h_intro hid; *) +(* Elim.h_decompose_and (mkVar hid); *) + tclTRY(list_rewrite true eqs); +(* (fun g -> *) +(* let ids' = pf_ids_of_hyps g in *) +(* let ids = List.filter (fun id -> not (List.mem id ids)) ids' in *) +(* rewrite *) +(* ) *) + Eauto.gen_eauto false (false,5) [] (Some []) + ] + gls + + +let backtrack_eqs_until_hrec hrec eqs : tactic = + fun gls -> + let eqs = List.map mkVar eqs in + let rewrite = + tclFIRST (List.map Equality.rewriteRL eqs ) + in + let _,hrec_concl = decompose_prod (pf_type_of gls (mkVar hrec)) in + let f_app = array_last (snd (destApp hrec_concl)) in + let f = (fst (destApp f_app)) in + let rec backtrack : tactic = + fun g -> + let f_app = array_last (snd (destApp (pf_concl g))) in + match kind_of_term f_app with + | App(f',_) when eq_constr f' f -> tclIDTAC g + | _ -> tclTHEN rewrite backtrack g + in + backtrack gls + + + +let build_clause eqs = + { + Tacexpr.onhyps = + Some (List.map + (fun id -> (Rawterm.all_occurrences_expr,id),InHyp) + eqs + ); + Tacexpr.concl_occs = Rawterm.no_occurrences_expr + } + +let rec rewrite_eqs_in_eqs eqs = + match eqs with + | [] -> tclIDTAC + | eq::eqs -> + + tclTHEN + (tclMAP + (fun id gl -> + observe_tac + (Format.sprintf "rewrite %s in %s " (string_of_id eq) (string_of_id id)) + (tclTRY (Equality.general_rewrite_in true all_occurrences (* dep proofs also: *) true id (mkVar eq) false)) + gl + ) + eqs + ) + (rewrite_eqs_in_eqs eqs) + +let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic = + fun gls -> + (tclTHENSEQ + [ + backtrack_eqs_until_hrec hrec eqs; + (* observe_tac ("new_prove_with_tcc ( applying "^(string_of_id hrec)^" )" ) *) + (tclTHENS (* We must have exactly ONE subgoal !*) + (apply (mkVar hrec)) + [ tclTHENSEQ + [ + keep (tcc_hyps@eqs); + apply (Lazy.force acc_inv); + (fun g -> + if is_mes + then + unfold_in_concl [(all_occurrences, evaluable_of_global_reference (delayed_force ltof_ref))] g + else tclIDTAC g + ); + observe_tac "rew_and_finish" + (tclTHENLIST + [tclTRY(Recdef.list_rewrite false (List.map mkVar eqs)); + observe_tac "rewrite_eqs_in_eqs" (rewrite_eqs_in_eqs eqs); + (observe_tac "finishing using" + ( + tclCOMPLETE( + Eauto.eauto_with_bases + false + (true,5) + [Lazy.force refl_equal] + [Auto.Hint_db.empty empty_transparent_state false] + ) + ) + ) + ] + ) + ] + ]) + ]) + gls + + +let is_valid_hypothesis predicates_name = + let predicates_name = List.fold_right Idset.add predicates_name Idset.empty in + let is_pte typ = + if isApp typ + then + let pte,_ = destApp typ in + if isVar pte + then Idset.mem (destVar pte) predicates_name + else false + else false + in + let rec is_valid_hypothesis typ = + is_pte typ || + match kind_of_term typ with + | Prod(_,pte,typ') -> is_pte pte && is_valid_hypothesis typ' + | _ -> false + in + is_valid_hypothesis + +let prove_principle_for_gen + (f_ref,functional_ref,eq_ref) tcc_lemma_ref is_mes + rec_arg_num rec_arg_type relation gl = + let princ_type = pf_concl gl in + let princ_info = compute_elim_sig princ_type in + let fresh_id = + let avoid = ref (pf_ids_of_hyps gl) in + fun na -> + let new_id = + match na with + | Name id -> fresh_id !avoid (string_of_id id) + | Anonymous -> fresh_id !avoid "H" + in + avoid := new_id :: !avoid; + Name new_id + in + let fresh_decl (na,b,t) = (fresh_id na,b,t) in + let princ_info : elim_scheme = + { princ_info with + params = List.map fresh_decl princ_info.params; + predicates = List.map fresh_decl princ_info.predicates; + branches = List.map fresh_decl princ_info.branches; + args = List.map fresh_decl princ_info.args + } + in + let wf_tac = + if is_mes + then + (fun b -> Recdef.tclUSER_if_not_mes tclIDTAC b None) + else fun _ -> prove_with_tcc tcc_lemma_ref [] + in + let real_rec_arg_num = rec_arg_num - princ_info.nparams in + let npost_rec_arg = princ_info.nargs - real_rec_arg_num + 1 in +(* observe ( *) +(* str "princ_type := " ++ pr_lconstr princ_type ++ fnl () ++ *) +(* str "princ_info.nparams := " ++ int princ_info.nparams ++ fnl () ++ *) + +(* str "princ_info.nargs := " ++ int princ_info.nargs ++ fnl () ++ *) +(* str "rec_arg_num := " ++ int rec_arg_num ++ fnl() ++ *) +(* str "real_rec_arg_num := " ++ int real_rec_arg_num ++ fnl () ++ *) +(* str "npost_rec_arg := " ++ int npost_rec_arg ); *) + let (post_rec_arg,pre_rec_arg) = + Util.list_chop npost_rec_arg princ_info.args + in + let rec_arg_id = + match List.rev post_rec_arg with + | (Name id,_,_)::_ -> id + | _ -> assert false + in +(* observe (str "rec_arg_id := " ++ pr_lconstr (mkVar rec_arg_id)); *) + let subst_constrs = List.map (fun (na,_,_) -> mkVar (Nameops.out_name na)) (pre_rec_arg@princ_info.params) in + let relation = substl subst_constrs relation in + let input_type = substl subst_constrs rec_arg_type in + let wf_thm_id = Nameops.out_name (fresh_id (Name (id_of_string "wf_R"))) in + let acc_rec_arg_id = + Nameops.out_name (fresh_id (Name (id_of_string ("Acc_"^(string_of_id rec_arg_id))))) + in + let revert l = + tclTHEN (h_generalize (List.map mkVar l)) (clear l) + in + let fix_id = Nameops.out_name (fresh_id (Name hrec_id)) in + let prove_rec_arg_acc g = + ((* observe_tac "prove_rec_arg_acc" *) + (tclCOMPLETE + (tclTHEN + (assert_by (Name wf_thm_id) + (mkApp (delayed_force well_founded,[|input_type;relation|])) + (fun g -> (* observe_tac "prove wf" *) (tclCOMPLETE (wf_tac is_mes)) g)) + ( + (* observe_tac *) +(* "apply wf_thm" *) + h_simplest_apply (mkApp(mkVar wf_thm_id,[|mkVar rec_arg_id|])) + ) + ) + ) + ) + g + in + let args_ids = List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.args in + let lemma = + match !tcc_lemma_ref with + | None -> anomaly ( "No tcc proof !!") + | Some lemma -> lemma + in +(* let rec list_diff del_list check_list = *) +(* match del_list with *) +(* [] -> *) +(* [] *) +(* | f::r -> *) +(* if List.mem f check_list then *) +(* list_diff r check_list *) +(* else *) +(* f::(list_diff r check_list) *) +(* in *) + let tcc_list = ref [] in + let start_tac gls = + let hyps = pf_ids_of_hyps gls in + let hid = + next_ident_away_in_goal + (id_of_string "prov") + hyps + in + tclTHENSEQ + [ + generalize [lemma]; + h_intro hid; + Elim.h_decompose_and (mkVar hid); + (fun g -> + let new_hyps = pf_ids_of_hyps g in + tcc_list := List.rev (list_subtract new_hyps (hid::hyps)); + if !tcc_list = [] + then + begin + tcc_list := [hid]; + tclIDTAC g + end + else thin [hid] g + ) + ] + gls + in + tclTHENSEQ + [ + observe_tac "start_tac" start_tac; + h_intros + (List.rev_map (fun (na,_,_) -> Nameops.out_name na) + (princ_info.args@princ_info.branches@princ_info.predicates@princ_info.params) + ); + (* observe_tac "" *) (assert_by + (Name acc_rec_arg_id) + (mkApp (delayed_force acc_rel,[|input_type;relation;mkVar rec_arg_id|])) + (prove_rec_arg_acc) + ); +(* observe_tac "reverting" *) (revert (List.rev (acc_rec_arg_id::args_ids))); +(* (fun g -> observe (Printer.pr_goal (sig_it g) ++ fnl () ++ *) +(* str "fix arg num" ++ int (List.length args_ids + 1) ); tclIDTAC g); *) + (* observe_tac "h_fix " *) (h_fix (Some fix_id) (List.length args_ids + 1)); +(* (fun g -> observe (Printer.pr_goal (sig_it g) ++ fnl() ++ pr_lconstr_env (pf_env g ) (pf_type_of g (mkVar fix_id) )); tclIDTAC g); *) + h_intros (List.rev (acc_rec_arg_id::args_ids)); + Equality.rewriteLR (mkConst eq_ref); + (* observe_tac "finish" *) (fun gl' -> + let body = + let _,args = destApp (pf_concl gl') in + array_last args + in + let body_info rec_hyps = + { + nb_rec_hyps = List.length rec_hyps; + rec_hyps = rec_hyps; + eq_hyps = []; + info = body + } + in + let acc_inv = + lazy ( + mkApp ( + delayed_force acc_inv_id, + [|input_type;relation;mkVar rec_arg_id|] + ) + ) + in + let acc_inv = lazy (mkApp(Lazy.force acc_inv, [|mkVar acc_rec_arg_id|])) in + let predicates_names = + List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.predicates + in + let pte_info = + { proving_tac = + (fun eqs -> +(* msgnl (str "tcc_list := "++ prlist_with_sep spc Ppconstr.pr_id !tcc_list); *) +(* msgnl (str "princ_info.args := "++ prlist_with_sep spc Ppconstr.pr_id (List.map (fun (na,_,_) -> (Nameops.out_name na)) princ_info.args)); *) +(* msgnl (str "princ_info.params := "++ prlist_with_sep spc Ppconstr.pr_id (List.map (fun (na,_,_) -> (Nameops.out_name na)) princ_info.params)); *) +(* msgnl (str "acc_rec_arg_id := "++ Ppconstr.pr_id acc_rec_arg_id); *) +(* msgnl (str "eqs := "++ prlist_with_sep spc Ppconstr.pr_id eqs); *) + + (* observe_tac "new_prove_with_tcc" *) + (new_prove_with_tcc + is_mes acc_inv fix_id + + (!tcc_list@(List.map + (fun (na,_,_) -> (Nameops.out_name na)) + (princ_info.args@princ_info.params) + )@ ([acc_rec_arg_id])) eqs + ) + + ); + is_valid = is_valid_hypothesis predicates_names + } + in + let ptes_info : pte_info Idmap.t = + List.fold_left + (fun map pte_id -> + Idmap.add pte_id + pte_info + map + ) + Idmap.empty + predicates_names + in + let make_proof rec_hyps = + build_proof + false + [f_ref] + ptes_info + (body_info rec_hyps) + in + (* observe_tac "instanciate_hyps_with_args" *) + (instanciate_hyps_with_args + make_proof + (List.map (fun (na,_,_) -> Nameops.out_name na) princ_info.branches) + (List.rev args_ids) + ) + gl' + ) + + ] + gl + + + + + + + + diff --git a/plugins/funind/functional_principles_proofs.mli b/plugins/funind/functional_principles_proofs.mli new file mode 100644 index 00000000..ff98f2b9 --- /dev/null +++ b/plugins/funind/functional_principles_proofs.mli @@ -0,0 +1,19 @@ +open Names +open Term + +val prove_princ_for_struct : + bool -> + int -> constant array -> constr array -> int -> Tacmach.tactic + + +val prove_principle_for_gen : + constant*constant*constant -> (* name of the function, the fonctionnal and the fixpoint equation *) + constr option ref -> (* a pointer to the obligation proofs lemma *) + bool -> (* is that function uses measure *) + int -> (* the number of recursive argument *) + types -> (* the type of the recursive argument *) + constr -> (* the wf relation used to prove the function *) + Tacmach.tactic + + +(* val is_pte : rel_declaration -> bool *) diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml new file mode 100644 index 00000000..b756492b --- /dev/null +++ b/plugins/funind/functional_principles_types.ml @@ -0,0 +1,737 @@ +open Printer +open Util +open Term +open Termops +open Namegen +open Names +open Declarations +open Pp +open Entries +open Hiddentac +open Evd +open Tacmach +open Proof_type +open Tacticals +open Tactics +open Indfun_common +open Functional_principles_proofs + +exception Toberemoved_with_rel of int*constr +exception Toberemoved + + +let pr_elim_scheme el = + let env = Global.env () in + let msg = str "params := " ++ Printer.pr_rel_context env el.params in + let env = Environ.push_rel_context el.params env in + let msg = msg ++ fnl () ++ str "predicates := "++ Printer.pr_rel_context env el.predicates in + let env = Environ.push_rel_context el.predicates env in + let msg = msg ++ fnl () ++ str "branches := " ++ Printer.pr_rel_context env el.branches in + let env = Environ.push_rel_context el.branches env in + let msg = msg ++ fnl () ++ str "args := " ++ Printer.pr_rel_context env el.args in + let env = Environ.push_rel_context el.args env in + msg ++ fnl () ++ str "concl := " ++ pr_lconstr_env env el.concl + + +let observe s = + if do_observe () + then Pp.msgnl s + + +let pr_elim_scheme el = + let env = Global.env () in + let msg = str "params := " ++ Printer.pr_rel_context env el.params in + let env = Environ.push_rel_context el.params env in + let msg = msg ++ fnl () ++ str "predicates := "++ Printer.pr_rel_context env el.predicates in + let env = Environ.push_rel_context el.predicates env in + let msg = msg ++ fnl () ++ str "branches := " ++ Printer.pr_rel_context env el.branches in + let env = Environ.push_rel_context el.branches env in + let msg = msg ++ fnl () ++ str "args := " ++ Printer.pr_rel_context env el.args in + let env = Environ.push_rel_context el.args env in + msg ++ fnl () ++ str "concl := " ++ pr_lconstr_env env el.concl + + +let observe s = + if do_observe () + then Pp.msgnl s + +(* + Transform an inductive induction principle into + a functional one +*) +let compute_new_princ_type_from_rel rel_to_fun sorts princ_type = + let princ_type_info = compute_elim_sig princ_type in + let env = Global.env () in + let env_with_params = Environ.push_rel_context princ_type_info.params env in + let tbl = Hashtbl.create 792 in + let rec change_predicates_names (avoid:identifier list) (predicates:rel_context) : rel_context = + match predicates with + | [] -> [] + |(Name x,v,t)::predicates -> + let id = Namegen.next_ident_away x avoid in + Hashtbl.add tbl id x; + (Name id,v,t)::(change_predicates_names (id::avoid) predicates) + | (Anonymous,_,_)::_ -> anomaly "Anonymous property binder " + in + let avoid = (Termops.ids_of_context env_with_params ) in + let princ_type_info = + { princ_type_info with + predicates = change_predicates_names avoid princ_type_info.predicates + } + in +(* observe (str "starting princ_type := " ++ pr_lconstr_env env princ_type); *) +(* observe (str "princ_infos : " ++ pr_elim_scheme princ_type_info); *) + let change_predicate_sort i (x,_,t) = + let new_sort = sorts.(i) in + let args,_ = decompose_prod t in + let real_args = + if princ_type_info.indarg_in_concl + then List.tl args + else args + in + Nameops.out_name x,None,compose_prod real_args (mkSort new_sort) + in + let new_predicates = + list_map_i + change_predicate_sort + 0 + princ_type_info.predicates + in + let env_with_params_and_predicates = List.fold_right Environ.push_named new_predicates env_with_params in + let rel_as_kn = + fst (match princ_type_info.indref with + | Some (Libnames.IndRef ind) -> ind + | _ -> error "Not a valid predicate" + ) + in + let ptes_vars = List.map (fun (id,_,_) -> id) new_predicates in + let is_pte = + let set = List.fold_right Idset.add ptes_vars Idset.empty in + fun t -> + match kind_of_term t with + | Var id -> Idset.mem id set + | _ -> false + in + let pre_princ = + it_mkProd_or_LetIn + ~init: + (it_mkProd_or_LetIn + ~init:(Option.fold_right + mkProd_or_LetIn + princ_type_info.indarg + princ_type_info.concl + ) + princ_type_info.args + ) + princ_type_info.branches + in + let pre_princ = substl (List.map mkVar ptes_vars) pre_princ in + let is_dom c = + match kind_of_term c with + | Ind((u,_)) -> u = rel_as_kn + | Construct((u,_),_) -> u = rel_as_kn + | _ -> false + in + let get_fun_num c = + match kind_of_term c with + | Ind(_,num) -> num + | Construct((_,num),_) -> num + | _ -> assert false + in + let dummy_var = mkVar (id_of_string "________") in + let mk_replacement c i args = + let res = mkApp(rel_to_fun.(i),Array.map pop (array_get_start args)) in +(* observe (str "replacing " ++ pr_lconstr c ++ str " by " ++ pr_lconstr res); *) + res + in + let rec has_dummy_var t = + fold_constr + (fun b t -> b || (eq_constr t dummy_var) || (has_dummy_var t)) + false + t + in + let rec compute_new_princ_type remove env pre_princ : types*(constr list) = + let (new_princ_type,_) as res = + match kind_of_term pre_princ with + | Rel n -> + begin + try match Environ.lookup_rel n env with + | _,_,t when is_dom t -> raise Toberemoved + | _ -> pre_princ,[] with Not_found -> assert false + end + | Prod(x,t,b) -> + compute_new_princ_type_for_binder remove mkProd env x t b + | Lambda(x,t,b) -> + compute_new_princ_type_for_binder remove mkLambda env x t b + | Ind _ | Construct _ when is_dom pre_princ -> raise Toberemoved + | App(f,args) when is_dom f -> + let var_to_be_removed = destRel (array_last args) in + let num = get_fun_num f in + raise (Toberemoved_with_rel (var_to_be_removed,mk_replacement pre_princ num args)) + | App(f,args) -> + let args = + if is_pte f && remove + then array_get_start args + else args + in + let new_args,binders_to_remove = + Array.fold_right (compute_new_princ_type_with_acc remove env) + args + ([],[]) + in + let new_f,binders_to_remove_from_f = compute_new_princ_type remove env f in + applist(new_f, new_args), + list_union_eq eq_constr binders_to_remove_from_f binders_to_remove + | LetIn(x,v,t,b) -> + compute_new_princ_type_for_letin remove env x v t b + | _ -> pre_princ,[] + in +(* let _ = match kind_of_term pre_princ with *) +(* | Prod _ -> *) +(* observe(str "compute_new_princ_type for "++ *) +(* pr_lconstr_env env pre_princ ++ *) +(* str" is "++ *) +(* pr_lconstr_env env new_princ_type ++ fnl ()) *) +(* | _ -> () in *) + res + + and compute_new_princ_type_for_binder remove bind_fun env x t b = + begin + try + let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in + let new_x : name = get_name (ids_of_context env) x in + let new_env = Environ.push_rel (x,None,t) env in + let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in + if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b + then (pop new_b),filter_map (eq_constr (mkRel 1)) pop binders_to_remove_from_b + else + ( + bind_fun(new_x,new_t,new_b), + list_union_eq + eq_constr + binders_to_remove_from_t + (List.map pop binders_to_remove_from_b) + ) + + with + | Toberemoved -> +(* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) + let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in + new_b, List.map pop binders_to_remove_from_b + | Toberemoved_with_rel (n,c) -> +(* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) + let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in + new_b, list_add_set_eq eq_constr (mkRel n) (List.map pop binders_to_remove_from_b) + end + and compute_new_princ_type_for_letin remove env x v t b = + begin + try + let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in + let new_v,binders_to_remove_from_v = compute_new_princ_type remove env v in + let new_x : name = get_name (ids_of_context env) x in + let new_env = Environ.push_rel (x,Some v,t) env in + let new_b,binders_to_remove_from_b = compute_new_princ_type remove new_env b in + if List.exists (eq_constr (mkRel 1)) binders_to_remove_from_b + then (pop new_b),filter_map (eq_constr (mkRel 1)) pop binders_to_remove_from_b + else + ( + mkLetIn(new_x,new_v,new_t,new_b), + list_union_eq + eq_constr + (list_union_eq eq_constr binders_to_remove_from_t binders_to_remove_from_v) + (List.map pop binders_to_remove_from_b) + ) + + with + | Toberemoved -> +(* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) + let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [dummy_var] 1 b) in + new_b, List.map pop binders_to_remove_from_b + | Toberemoved_with_rel (n,c) -> +(* observe (str "Decl of "++Ppconstr.pr_name x ++ str " is removed "); *) + let new_b,binders_to_remove_from_b = compute_new_princ_type remove env (substnl [c] n b) in + new_b, list_add_set_eq eq_constr (mkRel n) (List.map pop binders_to_remove_from_b) + end + and compute_new_princ_type_with_acc remove env e (c_acc,to_remove_acc) = + let new_e,to_remove_from_e = compute_new_princ_type remove env e + in + new_e::c_acc,list_union_eq eq_constr to_remove_from_e to_remove_acc + in +(* observe (str "Computing new principe from " ++ pr_lconstr_env env_with_params_and_predicates pre_princ); *) + let pre_res,_ = + compute_new_princ_type princ_type_info.indarg_in_concl env_with_params_and_predicates pre_princ + in + let pre_res = + replace_vars + (list_map_i (fun i id -> (id, mkRel i)) 1 ptes_vars) + (lift (List.length ptes_vars) pre_res) + in + it_mkProd_or_LetIn + ~init:(it_mkProd_or_LetIn + ~init:pre_res (List.map (fun (id,t,b) -> Name(Hashtbl.find tbl id), t,b) + new_predicates) + ) + princ_type_info.params + + + +let change_property_sort toSort princ princName = + let princ_info = compute_elim_sig princ in + let change_sort_in_predicate (x,v,t) = + (x,None, + let args,_ = decompose_prod t in + compose_prod args (mkSort toSort) + ) + in + let princName_as_constr = Tacinterp.constr_of_id (Global.env ()) princName in + let init = + let nargs = (princ_info.nparams + (List.length princ_info.predicates)) in + mkApp(princName_as_constr, + Array.init nargs + (fun i -> mkRel (nargs - i ))) + in + it_mkLambda_or_LetIn + ~init: + (it_mkLambda_or_LetIn ~init + (List.map change_sort_in_predicate princ_info.predicates) + ) + princ_info.params + + +let pp_dur time time' = + str (string_of_float (System.time_difference time time')) + +(* let qed () = save_named true *) +let defined () = + try + Lemmas.save_named false + with + | UserError("extract_proof",msg) -> + Util.errorlabstrm + "defined" + ((try + str "On goal : " ++ fnl () ++ pr_open_subgoals () ++ fnl () + with _ -> mt () + ) ++msg) + | e -> raise e + + + +let build_functional_principle interactive_proof old_princ_type sorts funs i proof_tac hook = + (* First we get the type of the old graph principle *) + let mutr_nparams = (compute_elim_sig old_princ_type).nparams in + (* let time1 = System.get_time () in *) + let new_principle_type = + compute_new_princ_type_from_rel + (Array.map mkConst funs) + sorts + old_princ_type + in + (* let time2 = System.get_time () in *) + (* Pp.msgnl (str "computing principle type := " ++ System.fmt_time_difference time1 time2); *) + observe (str "new_principle_type : " ++ pr_lconstr new_principle_type); + let new_princ_name = + next_ident_away_in_goal (id_of_string "___________princ_________") [] + in + begin + Lemmas.start_proof + new_princ_name + (Decl_kinds.Global,(Decl_kinds.Proof Decl_kinds.Theorem)) + new_principle_type + (hook new_principle_type) + ; + (* let _tim1 = System.get_time () in *) + Pfedit.by (proof_tac (Array.map mkConst funs) mutr_nparams); + (* let _tim2 = System.get_time () in *) + (* begin *) + (* let dur1 = System.time_difference tim1 tim2 in *) + (* Pp.msgnl (str ("Time to compute proof: ") ++ str (string_of_float dur1)); *) + (* end; *) + get_proof_clean true + end + + + +let generate_functional_principle + interactive_proof + old_princ_type sorts new_princ_name funs i proof_tac + = + try + + let f = funs.(i) in + let type_sort = Termops.new_sort_in_family InType in + let new_sorts = + match sorts with + | None -> Array.make (Array.length funs) (type_sort) + | Some a -> a + in + let base_new_princ_name,new_princ_name = + match new_princ_name with + | Some (id) -> id,id + | None -> + let id_of_f = id_of_label (con_label f) in + id_of_f,Indrec.make_elimination_ident id_of_f (family_of_sort type_sort) + in + let names = ref [new_princ_name] in + let hook new_principle_type _ _ = + if sorts = None + then + (* let id_of_f = id_of_label (con_label f) in *) + let register_with_sort fam_sort = + let s = Termops.new_sort_in_family fam_sort in + let name = Indrec.make_elimination_ident base_new_princ_name fam_sort in + let value = change_property_sort s new_principle_type new_princ_name in + (* Pp.msgnl (str "new principle := " ++ pr_lconstr value); *) + let ce = + { const_entry_body = value; + const_entry_type = None; + const_entry_opaque = false; + const_entry_boxed = Flags.boxed_definitions() + } + in + ignore( + Declare.declare_constant + name + (Entries.DefinitionEntry ce, + Decl_kinds.IsDefinition (Decl_kinds.Scheme) + ) + ); + Flags.if_verbose + (fun id -> Pp.msgnl (Ppconstr.pr_id id ++ str " is defined")) + name; + names := name :: !names + in + register_with_sort InProp; + register_with_sort InSet + in + let (id,(entry,g_kind,hook)) = + build_functional_principle interactive_proof old_princ_type new_sorts funs i proof_tac hook + in + (* Pr 1278 : + Don't forget to close the goal if an error is raised !!!! + *) + save false new_princ_name entry g_kind hook + with e -> + begin + begin + try + let id = Pfedit.get_current_proof_name () in + let s = string_of_id id in + let n = String.length "___________princ_________" in + if String.length s >= n + then if String.sub s 0 n = "___________princ_________" + then Pfedit.delete_current_proof () + else () + else () + with _ -> () + end; + raise (Defining_principle e) + end +(* defined () *) + + +exception Not_Rec + +let get_funs_constant mp dp = + let rec get_funs_constant const e : (Names.constant*int) array = + match kind_of_term ((strip_lam e)) with + | Fix((_,(na,_,_))) -> + Array.mapi + (fun i na -> + match na with + | Name id -> + let const = make_con mp dp (label_of_id id) in + const,i + | Anonymous -> + anomaly "Anonymous fix" + ) + na + | _ -> [|const,0|] + in + function const -> + let find_constant_body const = + match (Global.lookup_constant const ).const_body with + | Some b -> + let body = force b in + let body = Tacred.cbv_norm_flags + (Closure.RedFlags.mkflags [Closure.RedFlags.fZETA]) + (Global.env ()) + (Evd.empty) + body + in + body + | None -> error ( "Cannot define a principle over an axiom ") + in + let f = find_constant_body const in + let l_const = get_funs_constant const f in + (* + We need to check that all the functions found are in the same block + to prevent Reset stange thing + *) + let l_bodies = List.map find_constant_body (Array.to_list (Array.map fst l_const)) in + let l_params,l_fixes = List.split (List.map decompose_lam l_bodies) in + (* all the paremeter must be equal*) + let _check_params = + let first_params = List.hd l_params in + List.iter + (fun params -> + if not ((=) first_params params) + then error "Not a mutal recursive block" + ) + l_params + in + (* The bodies has to be very similar *) + let _check_bodies = + try + let extract_info is_first body = + match kind_of_term body with + | Fix((idxs,_),(na,ta,ca)) -> (idxs,na,ta,ca) + | _ -> + if is_first && (List.length l_bodies = 1) + then raise Not_Rec + else error "Not a mutal recursive block" + in + let first_infos = extract_info true (List.hd l_bodies) in + let check body = (* Hope this is correct *) + if not (first_infos = (extract_info false body)) + then error "Not a mutal recursive block" + in + List.iter check l_bodies + with Not_Rec -> () + in + l_const + +exception No_graph_found +exception Found_type of int + +let make_scheme (fas : (constant*Rawterm.rawsort) list) : Entries.definition_entry list = + let env = Global.env () + and sigma = Evd.empty in + let funs = List.map fst fas in + let first_fun = List.hd funs in + + + let funs_mp,funs_dp,_ = Names.repr_con first_fun in + let first_fun_kn = + try + fst (find_Function_infos first_fun).graph_ind + with Not_found -> raise No_graph_found + in + let this_block_funs_indexes = get_funs_constant funs_mp funs_dp first_fun in + let this_block_funs = Array.map fst this_block_funs_indexes in + let prop_sort = InProp in + let funs_indexes = + let this_block_funs_indexes = Array.to_list this_block_funs_indexes in + List.map + (function const -> List.assoc const this_block_funs_indexes) + funs + in + let ind_list = + List.map + (fun (idx) -> + let ind = first_fun_kn,idx in + ind,true,prop_sort + ) + funs_indexes + in + let l_schemes = + List.map + (Typing.type_of env sigma) + (Indrec.build_mutual_induction_scheme env sigma ind_list) + in + let i = ref (-1) in + let sorts = + List.rev_map (fun (_,x) -> + Termops.new_sort_in_family (Pretyping.interp_elimination_sort x) + ) + fas + in + (* We create the first priciple by tactic *) + let first_type,other_princ_types = + match l_schemes with + s::l_schemes -> s,l_schemes + | _ -> anomaly "" + in + let (_,(const,_,_)) = + try + build_functional_principle false + first_type + (Array.of_list sorts) + this_block_funs + 0 + (prove_princ_for_struct false 0 (Array.of_list funs)) + (fun _ _ _ -> ()) + with e -> + begin + begin + try + let id = Pfedit.get_current_proof_name () in + let s = string_of_id id in + let n = String.length "___________princ_________" in + if String.length s >= n + then if String.sub s 0 n = "___________princ_________" + then Pfedit.delete_current_proof () + else () + else () + with _ -> () + end; + raise (Defining_principle e) + end + + in + incr i; + let opacity = + let finfos = find_Function_infos this_block_funs.(0) in + try + let equation = Option.get finfos.equation_lemma in + (Global.lookup_constant equation).Declarations.const_opaque + with Option.IsNone -> (* non recursive definition *) + false + in + let const = {const with const_entry_opaque = opacity } in + (* The others are just deduced *) + if other_princ_types = [] + then + [const] + else + let other_fun_princ_types = + let funs = Array.map mkConst this_block_funs in + let sorts = Array.of_list sorts in + List.map (compute_new_princ_type_from_rel funs sorts) other_princ_types + in + let first_princ_body,first_princ_type = const.Entries.const_entry_body, const.Entries.const_entry_type in + let ctxt,fix = decompose_lam_assum first_princ_body in (* the principle has for forall ...., fix .*) + let (idxs,_),(_,ta,_ as decl) = destFix fix in + let other_result = + List.map (* we can now compute the other principles *) + (fun scheme_type -> + incr i; + observe (Printer.pr_lconstr scheme_type); + let type_concl = (strip_prod_assum scheme_type) in + let applied_f = List.hd (List.rev (snd (decompose_app type_concl))) in + let f = fst (decompose_app applied_f) in + try (* we search the number of the function in the fix block (name of the function) *) + Array.iteri + (fun j t -> + let t = (strip_prod_assum t) in + let applied_g = List.hd (List.rev (snd (decompose_app t))) in + let g = fst (decompose_app applied_g) in + if eq_constr f g + then raise (Found_type j); + observe (Printer.pr_lconstr f ++ str " <> " ++ + Printer.pr_lconstr g) + + ) + ta; + (* If we reach this point, the two principle are not mutually recursive + We fall back to the previous method + *) + let (_,(const,_,_)) = + build_functional_principle + false + (List.nth other_princ_types (!i - 1)) + (Array.of_list sorts) + this_block_funs + !i + (prove_princ_for_struct false !i (Array.of_list funs)) + (fun _ _ _ -> ()) + in + const + with Found_type i -> + let princ_body = + Termops.it_mkLambda_or_LetIn ~init:(mkFix((idxs,i),decl)) ctxt + in + {const with + Entries.const_entry_body = princ_body; + Entries.const_entry_type = Some scheme_type + } + ) + other_fun_princ_types + in + const::other_result + +let build_scheme fas = + Dumpglob.pause (); + let bodies_types = + make_scheme + (List.map + (fun (_,f,sort) -> + let f_as_constant = + try + match Nametab.global f with + | Libnames.ConstRef c -> c + | _ -> Util.error "Functional Scheme can only be used with functions" + with Not_found -> + Util.error ("Cannot find "^ Libnames.string_of_reference f) + in + (f_as_constant,sort) + ) + fas + ) + in + List.iter2 + (fun (princ_id,_,_) def_entry -> + ignore + (Declare.declare_constant + princ_id + (Entries.DefinitionEntry def_entry,Decl_kinds.IsProof Decl_kinds.Theorem)); + Flags.if_verbose + (fun id -> Pp.msgnl (Ppconstr.pr_id id ++ str " is defined")) princ_id + ) + fas + bodies_types; + Dumpglob.continue () + + + +let build_case_scheme fa = + let env = Global.env () + and sigma = Evd.empty in +(* let id_to_constr id = *) +(* Tacinterp.constr_of_id env id *) +(* in *) + let funs = (fun (_,f,_) -> + try Libnames.constr_of_global (Nametab.global f) + with Not_found -> + Util.error ("Cannot find "^ Libnames.string_of_reference f)) fa in + let first_fun = destConst funs in + + let funs_mp,funs_dp,_ = Names.repr_con first_fun in + let first_fun_kn = try fst (find_Function_infos first_fun).graph_ind with Not_found -> raise No_graph_found in + + + + let this_block_funs_indexes = get_funs_constant funs_mp funs_dp first_fun in + let this_block_funs = Array.map fst this_block_funs_indexes in + let prop_sort = InProp in + let funs_indexes = + let this_block_funs_indexes = Array.to_list this_block_funs_indexes in + List.assoc (destConst funs) this_block_funs_indexes + in + let ind_fun = + let ind = first_fun_kn,funs_indexes in + ind,prop_sort + in + let scheme_type = (Typing.type_of env sigma ) ((fun (ind,sf) -> Indrec.build_case_analysis_scheme_default env sigma ind sf) ind_fun) in + let sorts = + (fun (_,_,x) -> + Termops.new_sort_in_family (Pretyping.interp_elimination_sort x) + ) + fa + in + let princ_name = (fun (x,_,_) -> x) fa in + let _ = + (* Pp.msgnl (str "Generating " ++ Ppconstr.pr_id princ_name ++str " with " ++ + pr_lconstr scheme_type ++ str " and " ++ (fun a -> prlist_with_sep spc (fun c -> pr_lconstr (mkConst c)) (Array.to_list a)) this_block_funs + ); + *) + generate_functional_principle + false + scheme_type + (Some ([|sorts|])) + (Some princ_name) + this_block_funs + 0 + (prove_princ_for_struct false 0 [|destConst funs|]) + in + () diff --git a/plugins/funind/functional_principles_types.mli b/plugins/funind/functional_principles_types.mli new file mode 100644 index 00000000..fb04c6ec --- /dev/null +++ b/plugins/funind/functional_principles_types.mli @@ -0,0 +1,34 @@ +open Names +open Term + + +val generate_functional_principle : + (* do we accept interactive proving *) + bool -> + (* induction principle on rel *) + types -> + (* *) + sorts array option -> + (* Name of the new principle *) + (identifier) option -> + (* the compute functions to use *) + constant array -> + (* We prove the nth- principle *) + int -> + (* The tactic to use to make the proof w.r + the number of params + *) + (constr array -> int -> Tacmach.tactic) -> + unit + +val compute_new_princ_type_from_rel : constr array -> sorts array -> + types -> types + + +exception No_graph_found + +val make_scheme : (constant*Rawterm.rawsort) list -> Entries.definition_entry list + +val build_scheme : (identifier*Libnames.reference*Rawterm.rawsort) list -> unit +val build_case_scheme : (identifier*Libnames.reference*Rawterm.rawsort) -> unit + diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4 new file mode 100644 index 00000000..bc400ae1 --- /dev/null +++ b/plugins/funind/g_indfun.ml4 @@ -0,0 +1,524 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* hov 1 (Ppconstr.pr_id id ++ str " := " ++ cut () ++ prc c) + | loc, Rawterm.AnonHyp n, c -> hov 1 (int n ++ str " := " ++ cut () ++ prc c) + +let pr_bindings prc prlc = function + | Rawterm.ImplicitBindings l -> + brk (1,1) ++ str "with" ++ brk (1,1) ++ + Util.prlist_with_sep spc prc l + | Rawterm.ExplicitBindings l -> + brk (1,1) ++ str "with" ++ brk (1,1) ++ + Util.prlist_with_sep spc (fun b -> str"(" ++ pr_binding prlc b ++ str")") l + | Rawterm.NoBindings -> mt () + +let pr_with_bindings prc prlc (c,bl) = + prc c ++ hv 0 (pr_bindings prc prlc bl) + +let pr_fun_ind_using prc prlc _ opt_c = + match opt_c with + | None -> mt () + | Some b -> spc () ++ hov 2 (str "using" ++ spc () ++ pr_with_bindings prc prlc b) + +(* Duplication of printing functions because "'a with_bindings" is + (internally) not uniform in 'a: indeed constr_with_bindings at the + "typed" level has type "open_constr with_bindings" instead of + "constr with_bindings"; hence, its printer cannot be polymorphic in + (prc,prlc)... *) + +let pr_with_bindings_typed prc prlc (c,bl) = + prc c ++ + hv 0 (pr_bindings prc prlc bl) + +let pr_fun_ind_using_typed prc prlc _ opt_c = + match opt_c with + | None -> mt () + | Some b -> spc () ++ hov 2 (str "using" ++ spc () ++ pr_with_bindings_typed prc prlc b.Evd.it) + + +ARGUMENT EXTEND fun_ind_using + TYPED AS constr_with_bindings_opt + PRINTED BY pr_fun_ind_using_typed + RAW_TYPED AS constr_with_bindings_opt + RAW_PRINTED BY pr_fun_ind_using + GLOB_TYPED AS constr_with_bindings_opt + GLOB_PRINTED BY pr_fun_ind_using +| [ "using" constr_with_bindings(c) ] -> [ Some c ] +| [ ] -> [ None ] +END + + +TACTIC EXTEND newfuninv + [ "functional" "inversion" quantified_hypothesis(hyp) reference_opt(fname) ] -> + [ + Invfun.invfun hyp fname + ] +END + + +let pr_intro_as_pat prc _ _ pat = + match pat with + | Some pat -> spc () ++ str "as" ++ spc () ++ pr_intro_pattern pat + | None -> mt () + + +ARGUMENT EXTEND with_names TYPED AS intro_pattern_opt PRINTED BY pr_intro_as_pat +| [ "as" simple_intropattern(ipat) ] -> [ Some ipat ] +| [] ->[ None ] +END + + + + +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 -> applist(c,cl) + in + Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl ] +END +(***** debug only ***) +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 -> applist(c,cl) + in + Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl ] +END + + +let pr_constr_coma_sequence prc _ _ = Util.prlist_with_sep Util.pr_comma prc + +ARGUMENT EXTEND constr_coma_sequence' + TYPED AS constr_list + PRINTED BY pr_constr_coma_sequence +| [ constr(c) "," constr_coma_sequence'(l) ] -> [ c::l ] +| [ constr(c) ] -> [ [c] ] +END + +let pr_auto_using prc _prlc _prt = Pptactic.pr_auto_using prc + +ARGUMENT EXTEND auto_using' + TYPED AS constr_list + PRINTED BY pr_auto_using +| [ "using" constr_coma_sequence'(l) ] -> [ l ] +| [ ] -> [ [] ] +END + +let pr_rec_annotation2_aux s r id l = + str ("{"^s^" ") ++ Ppconstr.pr_constr_expr r ++ + Util.pr_opt Nameops.pr_id id ++ + Pptactic.pr_auto_using Ppconstr.pr_constr_expr l ++ str "}" + +let pr_rec_annotation2 = function + | Struct id -> str "{struct" ++ Nameops.pr_id id ++ str "}" + | Wf(r,id,l) -> pr_rec_annotation2_aux "wf" r id l + | Mes(r,id,l) -> pr_rec_annotation2_aux "measure" r id l + +VERNAC ARGUMENT EXTEND rec_annotation2 +PRINTED BY pr_rec_annotation2 + [ "{" "struct" ident(id) "}"] -> [ Struct id ] +| [ "{" "wf" constr(r) ident_opt(id) auto_using'(l) "}" ] -> [ Wf(r,id,l) ] +| [ "{" "measure" constr(r) ident_opt(id) auto_using'(l) "}" ] -> [ Mes(r,id,l) ] +END + +let pr_binder2 (idl,c) = + str "(" ++ Util.prlist_with_sep spc Nameops.pr_id idl ++ spc () ++ + str ": " ++ Ppconstr.pr_lconstr_expr c ++ str ")" + +VERNAC ARGUMENT EXTEND binder2 +PRINTED BY pr_binder2 + [ "(" ne_ident_list(idl) ":" lconstr(c) ")"] -> [ (idl,c) ] +END + +let make_binder2 (idl,c) = + LocalRawAssum (List.map (fun id -> (Util.dummy_loc,Name id)) idl,Topconstr.default_binder_kind,c) + +let pr_rec_definition2 (id,bl,annot,type_,def) = + Nameops.pr_id id ++ spc () ++ Util.prlist_with_sep spc pr_binder2 bl ++ + Util.pr_opt pr_rec_annotation2 annot ++ spc () ++ str ":" ++ spc () ++ + Ppconstr.pr_lconstr_expr type_ ++ str " :=" ++ spc () ++ + Ppconstr.pr_lconstr_expr def + +VERNAC ARGUMENT EXTEND rec_definition2 +PRINTED BY pr_rec_definition2 + [ ident(id) binder2_list(bl) + rec_annotation2_opt(annot) ":" lconstr(type_) + ":=" lconstr(def)] -> + [ (id,bl,annot,type_,def) ] +END + +let make_rec_definitions2 (id,bl,annot,type_,def) = + let bl = List.map make_binder2 bl in + let names = List.map snd (Topconstr.names_of_local_assums bl) in + let check_one_name () = + if List.length names > 1 then + Util.user_err_loc + (Util.dummy_loc,"Function", + Pp.str "the recursive argument needs to be specified"); + in + let check_exists_args an = + try + let id = match an with + | Struct id -> id | Wf(_,Some id,_) -> id | Mes(_,Some id,_) -> id + | Wf(_,None,_) | Mes(_,None,_) -> failwith "check_exists_args" + in + (try ignore(Util.list_index0 (Name id) names); annot + with Not_found -> Util.user_err_loc + (Util.dummy_loc,"Function", + Pp.str "No argument named " ++ Nameops.pr_id id) + ) + with Failure "check_exists_args" -> check_one_name ();annot + in + let ni = + match annot with + | None -> + annot + | Some an -> + check_exists_args an + in + ((Util.dummy_loc,id), ni, bl, type_, def) + + +VERNAC COMMAND EXTEND Function + ["Function" ne_rec_definition2_list_sep(recsl,"with")] -> + [ + do_generate_principle false (List.map make_rec_definitions2 recsl); + + ] +END + +let pr_fun_scheme_arg (princ_name,fun_name,s) = + Nameops.pr_id princ_name ++ str " :=" ++ spc() ++ str "Induction for " ++ + Libnames.pr_reference fun_name ++ spc() ++ str "Sort " ++ + Ppconstr.pr_rawsort s + +VERNAC ARGUMENT EXTEND fun_scheme_arg +PRINTED BY pr_fun_scheme_arg +| [ ident(princ_name) ":=" "Induction" "for" reference(fun_name) "Sort" sort(s) ] -> [ (princ_name,fun_name,s) ] +END + + +let warning_error names e = + match e with + | Building_graph e -> + Pp.msg_warning + (str "Cannot define graph(s) for " ++ + h 1 (prlist_with_sep (fun _ -> str","++spc ()) Libnames.pr_reference names) ++ + if do_observe () then (spc () ++ Cerrors.explain_exn e) else mt ()) + | Defining_principle e -> + Pp.msg_warning + (str "Cannot define principle(s) for "++ + h 1 (prlist_with_sep (fun _ -> str","++spc ()) Libnames.pr_reference names) ++ + if do_observe () then Cerrors.explain_exn e else mt ()) + | _ -> anomaly "" + + +VERNAC COMMAND EXTEND NewFunctionalScheme + ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ] -> + [ + begin + try + Functional_principles_types.build_scheme fas + with Functional_principles_types.No_graph_found -> + begin + match fas with + | (_,fun_name,_)::_ -> + begin + begin + make_graph (Nametab.global fun_name) + end + ; + try Functional_principles_types.build_scheme fas + with Functional_principles_types.No_graph_found -> + Util.error ("Cannot generate induction principle(s)") + | e -> + let names = List.map (fun (_,na,_) -> na) fas in + warning_error names e + + end + | _ -> assert false (* we can only have non empty list *) + end + | e -> + let names = List.map (fun (_,na,_) -> na) fas in + warning_error names e + + end + ] +END +(***** debug only ***) + +VERNAC COMMAND EXTEND NewFunctionalCase + ["Functional" "Case" fun_scheme_arg(fas) ] -> + [ + Functional_principles_types.build_case_scheme fas + ] +END + +(***** debug only ***) +VERNAC COMMAND EXTEND GenerateGraph +["Generate" "graph" "for" reference(c)] -> [ make_graph (Nametab.global c) ] +END + + + + + +(* FINDUCTION *) + +(* comment this line to see debug msgs *) +let msg x = () ;; let pr_lconstr c = str "" + (* uncomment this to see debugging *) +let prconstr c = msg (str" " ++ Printer.pr_lconstr c ++ str"\n") +let prlistconstr lc = List.iter prconstr lc +let prstr s = msg(str s) +let prNamedConstr s c = + begin + msg(str ""); + msg(str(s^"==>\n ") ++ Printer.pr_lconstr c ++ str "\n<==\n"); + msg(str ""); + end + + + +(** Information about an occurrence of a function call (application) + inside a term. *) +type fapp_info = { + fname: constr; (** The function applied *) + largs: constr list; (** List of arguments *) + free: bool; (** [true] if all arguments are debruijn free *) + max_rel: int; (** max debruijn index in the funcall *) + onlyvars: bool (** [true] if all arguments are variables (and not debruijn) *) +} + + +(** [constr_head_match(a b c) a] returns true, false otherwise. *) +let constr_head_match u t= + if isApp u + then + let uhd,args= destApp u in + uhd=t + else false + +(** [hdMatchSub inu t] returns the list of occurrences of [t] in + [inu]. DeBruijn are not pushed, so some of them may be unbound in + the result. *) +let rec hdMatchSub inu (test: constr -> bool) : fapp_info list = + let subres = + match kind_of_term inu with + | Lambda (nm,tp,cstr) | Prod (nm,tp,cstr) -> + hdMatchSub tp test @ hdMatchSub (lift 1 cstr) test + | Fix (_,(lna,tl,bl)) -> (* not sure Fix is correct *) + Array.fold_left + (fun acc cstr -> acc @ hdMatchSub (lift (Array.length tl) cstr) test) + [] bl + | _ -> (* Cofix will be wrong *) + fold_constr + (fun l cstr -> + l @ hdMatchSub cstr test) [] inu in + if not (test inu) then subres + else + let f,args = decompose_app inu in + let freeset = Termops.free_rels inu in + let max_rel = try Util.Intset.max_elt freeset with Not_found -> -1 in + {fname = f; largs = args; free = Util.Intset.is_empty freeset; + max_rel = max_rel; onlyvars = List.for_all isVar args } + ::subres + +let mkEq typ c1 c2 = + mkApp (Coqlib.build_coq_eq(),[| typ; c1; c2|]) + + +let poseq_unsafe idunsafe cstr gl = + let typ = Tacmach.pf_type_of gl cstr in + tclTHEN + (Tactics.letin_tac None (Name idunsafe) cstr None allHypsAndConcl) + (tclTHENFIRST + (Tactics.assert_tac Anonymous (mkEq typ (mkVar idunsafe) cstr)) + Tactics.reflexivity) + gl + + +let poseq id cstr gl = + let x = Tactics.fresh_id [] id gl in + poseq_unsafe x cstr gl + +(* dirty? *) + +let list_constr_largs = ref [] + +let rec poseq_list_ids_rec lcstr gl = + match lcstr with + | [] -> tclIDTAC gl + | c::lcstr' -> + match kind_of_term c with + | Var _ -> + (list_constr_largs:=c::!list_constr_largs ; poseq_list_ids_rec lcstr' gl) + | _ -> + let _ = prstr "c = " in + let _ = prconstr c in + let _ = prstr "\n" in + let typ = Tacmach.pf_type_of gl c in + let cname = Namegen.id_of_name_using_hdchar (Global.env()) typ Anonymous in + let x = Tactics.fresh_id [] cname gl in + let _ = list_constr_largs:=mkVar x :: !list_constr_largs in + let _ = prstr " list_constr_largs = " in + let _ = prlistconstr !list_constr_largs in + let _ = prstr "\n" in + + tclTHEN + (poseq_unsafe x c) + (poseq_list_ids_rec lcstr') + gl + +let poseq_list_ids lcstr gl = + let _ = list_constr_largs := [] in + poseq_list_ids_rec lcstr gl + +(** [find_fapp test g] returns the list of [app_info] of all calls to + functions that satisfy [test] in the conclusion of goal g. Trivial + repetition (not modulo conversion) are deleted. *) +let find_fapp (test:constr -> bool) g : fapp_info list = + let pre_res = hdMatchSub (Tacmach.pf_concl g) test in + let res = + List.fold_right (fun x acc -> if List.mem x acc then acc else x::acc) pre_res [] in + (prlistconstr (List.map (fun x -> applist (x.fname,x.largs)) res); + res) + + + +(** [finduction id filter g] tries to apply functional induction on + an occurence of function [id] in the conclusion of goal [g]. If + [id]=[None] then calls to any function are selected. In any case + [heuristic] is used to select the most pertinent occurrence. *) +let finduction (oid:identifier option) (heuristic: fapp_info list -> fapp_info list) + (nexttac:Proof_type.tactic) g = + let test = match oid with + | Some id -> + let idconstr = mkConst (const_of_id id) in + (fun u -> constr_head_match u idconstr) (* select only id *) + | None -> (fun u -> isApp u) in (* select calls to any function *) + let info_list = find_fapp test g in + let ordered_info_list = heuristic info_list in + prlistconstr (List.map (fun x -> applist (x.fname,x.largs)) ordered_info_list); + if List.length ordered_info_list = 0 then Util.error "function not found in goal\n"; + let taclist: Proof_type.tactic list = + List.map + (fun info -> + (tclTHEN + (tclTHEN (poseq_list_ids info.largs) + ( + fun gl -> + (functional_induction + true (applist (info.fname, List.rev !list_constr_largs)) + None None) gl)) + nexttac)) ordered_info_list in + (* we try each (f t u v) until one does not fail *) + (* TODO: try also to mix functional schemes *) + tclFIRST taclist g + + + + +(** [chose_heuristic oi x] returns the heuristic for reordering + (and/or forgetting some elts of) a list of occurrences of + function calls infos to chose first with functional induction. *) +let chose_heuristic (oi:int option) : fapp_info list -> fapp_info list = + match oi with + | Some i -> (fun l -> [ List.nth l (i-1) ]) (* occurrence was given by the user *) + | None -> + (* Default heuristic: put first occurrences where all arguments + are *bound* (meaning already introduced) variables *) + let ordering x y = + if x.free && x.onlyvars && y.free && y.onlyvars then 0 (* both pertinent *) + else if x.free && x.onlyvars then -1 + else if y.free && y.onlyvars then 1 + else 0 (* both not pertinent *) + in + List.sort ordering + + + +TACTIC EXTEND finduction + ["finduction" ident(id) natural_opt(oi)] -> + [ + match oi with + | Some(n) when n<=0 -> Util.error "numerical argument must be > 0" + | _ -> + let heuristic = chose_heuristic oi in + finduction (Some id) heuristic tclIDTAC + ] +END + + + +TACTIC EXTEND fauto + [ "fauto" tactic(tac)] -> + [ + let heuristic = chose_heuristic None in + finduction None heuristic (snd tac) + ] + | + [ "fauto" ] -> + [ + let heuristic = chose_heuristic None in + finduction None heuristic tclIDTAC + ] + +END + + +TACTIC EXTEND poseq + [ "poseq" ident(x) constr(c) ] -> + [ poseq x c ] +END + +VERNAC COMMAND EXTEND Showindinfo + [ "showindinfo" ident(x) ] -> [ Merge.showind x ] +END + +VERNAC COMMAND EXTEND MergeFunind + [ "Mergeschemes" "(" ident(id1) ne_ident_list(cl1) ")" + "with" "(" ident(id2) ne_ident_list(cl2) ")" "using" ident(id) ] -> + [ + let f1 = Constrintern.interp_constr Evd.empty (Global.env()) + (CRef (Libnames.Ident (Util.dummy_loc,id1))) in + let f2 = Constrintern.interp_constr Evd.empty (Global.env()) + (CRef (Libnames.Ident (Util.dummy_loc,id2))) in + let f1type = Typing.type_of (Global.env()) Evd.empty f1 in + let f2type = Typing.type_of (Global.env()) Evd.empty f2 in + let ar1 = List.length (fst (decompose_prod f1type)) in + let ar2 = List.length (fst (decompose_prod f2type)) in + let _ = + if ar1 <> List.length cl1 then + Util.error ("not the right number of arguments for " ^ string_of_id id1) in + let _ = + if ar2 <> List.length cl2 then + Util.error ("not the right number of arguments for " ^ string_of_id id2) in + Merge.merge id1 id2 (Array.of_list cl1) (Array.of_list cl2) id + ] +END diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml new file mode 100644 index 00000000..38f42844 --- /dev/null +++ b/plugins/funind/indfun.ml @@ -0,0 +1,776 @@ +open Util +open Names +open Term +open Pp +open Indfun_common +open Libnames +open Rawterm +open Declarations + +let is_rec_info scheme_info = + let test_branche min acc (_,_,br) = + acc || ( + let new_branche = + it_mkProd_or_LetIn mkProp (fst (decompose_prod_assum br)) in + let free_rels_in_br = Termops.free_rels new_branche in + let max = min + scheme_info.Tactics.npredicates in + Util.Intset.exists (fun i -> i >= min && i< max) free_rels_in_br + ) + in + Util.list_fold_left_i test_branche 1 false (List.rev scheme_info.Tactics.branches) + + +let choose_dest_or_ind scheme_info = + if is_rec_info scheme_info + then Tactics.new_induct false + else Tactics.new_destruct false + + +let functional_induction with_clean c princl pat = + Dumpglob.pause (); + let res = let f,args = decompose_app c in + fun g -> + let princ,bindings, princ_type = + match princl with + | None -> (* No principle is given let's find the good one *) + begin + match kind_of_term f with + | Const c' -> + let princ_option = + let finfo = (* we first try to find out a graph on f *) + try find_Function_infos c' + with Not_found -> + errorlabstrm "" (str "Cannot find induction information on "++ + Printer.pr_lconstr (mkConst c') ) + in + match Tacticals.elimination_sort_of_goal g with + | InProp -> finfo.prop_lemma + | InSet -> finfo.rec_lemma + | InType -> finfo.rect_lemma + in + let princ = (* then we get the principle *) + try mkConst (Option.get princ_option ) + with Option.IsNone -> + (*i If there is not default lemma defined then, + we cross our finger and try to find a lemma named f_ind + (or f_rec, f_rect) i*) + let princ_name = + Indrec.make_elimination_ident + (id_of_label (con_label c')) + (Tacticals.elimination_sort_of_goal g) + in + try + mkConst(const_of_id princ_name ) + with Not_found -> (* This one is neither defined ! *) + errorlabstrm "" (str "Cannot find induction principle for " + ++Printer.pr_lconstr (mkConst c') ) + in + (princ,Rawterm.NoBindings, Tacmach.pf_type_of g princ) + | _ -> raise (UserError("",str "functional induction must be used with a function" )) + + end + | Some ((princ,binding)) -> + princ,binding,Tacmach.pf_type_of g princ + in + let princ_infos = Tactics.compute_elim_sig princ_type in + let args_as_induction_constr = + let c_list = + if princ_infos.Tactics.farg_in_concl + then [c] else [] + in + List.map (fun c -> Tacexpr.ElimOnConstr (c,NoBindings)) (args@c_list) + in + let princ' = Some (princ,bindings) in + let princ_vars = + List.fold_right + (fun a acc -> + try Idset.add (destVar a) acc + with _ -> acc + ) + args + Idset.empty + in + let old_idl = List.fold_right Idset.add (Tacmach.pf_ids_of_hyps g) Idset.empty in + let old_idl = Idset.diff old_idl princ_vars in + let subst_and_reduce g = + if with_clean + then + let idl = + map_succeed + (fun id -> + if Idset.mem id old_idl then failwith "subst_and_reduce"; + id + ) + (Tacmach.pf_ids_of_hyps g) + in + let flag = + Rawterm.Cbv + {Rawterm.all_flags + with Rawterm.rDelta = false; + } + in + Tacticals.tclTHEN + (Tacticals.tclMAP (fun id -> Tacticals.tclTRY (Equality.subst_gen (do_rewrite_dependent ()) [id])) idl ) + (Hiddentac.h_reduce flag Tacticals.allHypsAndConcl) + g + else Tacticals.tclIDTAC g + + in + Tacticals.tclTHEN + (choose_dest_or_ind + princ_infos + args_as_induction_constr + princ' + (None,pat) + None) + subst_and_reduce + g + in + Dumpglob.continue (); + res + + + + +type annot = + Struct of identifier + | Wf of Topconstr.constr_expr * identifier option * Topconstr.constr_expr list + | Mes of Topconstr.constr_expr * identifier option * Topconstr.constr_expr list + + +type newfixpoint_expr = + identifier * annot * Topconstr.local_binder list * Topconstr.constr_expr * Topconstr.constr_expr + +let rec abstract_rawconstr c = function + | [] -> c + | Topconstr.LocalRawDef (x,b)::bl -> Topconstr.mkLetInC(x,b,abstract_rawconstr c bl) + | Topconstr.LocalRawAssum (idl,k,t)::bl -> + List.fold_right (fun x b -> Topconstr.mkLambdaC([x],k,t,b)) idl + (abstract_rawconstr c bl) + +let interp_casted_constr_with_implicits sigma env impls c = +(* Constrintern.interp_rawconstr_with_implicits sigma env [] impls c *) + Constrintern.intern_gen false sigma env ~impls + ~allow_patvar:false ~ltacvars:([],[]) c + + +(* + Construct a fixpoint as a Rawterm + and not as a constr +*) +let build_newrecursive +(lnameargsardef) = + let env0 = Global.env() + and sigma = Evd.empty + in + let (rec_sign,rec_impls) = + List.fold_left + (fun (env,impls) ((_,recname),_,bl,arityc,_) -> + let arityc = Topconstr.prod_constr_expr arityc bl in + let arity = Constrintern.interp_type sigma env0 arityc in + let impl = Constrintern.compute_internalization_data env0 Constrintern.Recursive arity [] in + (Environ.push_named (recname,None,arity) env, (recname,impl) :: impls)) + (env0,[]) lnameargsardef in + let rec_impls = Constrintern.set_internalization_env_params rec_impls [] in + let recdef = + (* Declare local notations *) + let fs = States.freeze() in + let def = + try + List.map + (fun (_,_,bl,_,def) -> + let def = abstract_rawconstr def bl in + interp_casted_constr_with_implicits + sigma rec_sign rec_impls def + ) + lnameargsardef + with e -> + States.unfreeze fs; raise e in + States.unfreeze fs; def + in + recdef,rec_impls + + +let compute_annot (name,annot,args,types,body) = + let names = List.map snd (Topconstr.names_of_local_assums args) in + match annot with + | None -> + if List.length names > 1 then + user_err_loc + (dummy_loc,"Function", + Pp.str "the recursive argument needs to be specified"); + let new_annot = (id_of_name (List.hd names)) in + (name,Struct new_annot,args,types,body) + | Some r -> (name,r,args,types,body) + + +(* Checks whether or not the mutual bloc is recursive *) +let rec is_rec names = + let names = List.fold_right Idset.add names Idset.empty in + let check_id id names = Idset.mem id names in + let rec lookup names = function + | RVar(_,id) -> check_id id names + | RRef _ | REvar _ | RPatVar _ | RSort _ | RHole _ | RDynamic _ -> false + | RCast(_,b,_) -> lookup names b + | RRec _ -> error "RRec not handled" + | RIf(_,b,_,lhs,rhs) -> + (lookup names b) || (lookup names lhs) || (lookup names rhs) + | RLetIn(_,na,t,b) | RLambda(_,na,_,t,b) | RProd(_,na,_,t,b) -> + lookup names t || lookup (Nameops.name_fold Idset.remove na names) b + | RLetTuple(_,nal,_,t,b) -> lookup names t || + lookup + (List.fold_left + (fun acc na -> Nameops.name_fold Idset.remove na acc) + names + nal + ) + b + | RApp(_,f,args) -> List.exists (lookup names) (f::args) + | RCases(_,_,_,el,brl) -> + List.exists (fun (e,_) -> lookup names e) el || + List.exists (lookup_br names) brl + and lookup_br names (_,idl,_,rt) = + let new_names = List.fold_right Idset.remove idl names in + lookup new_names rt + in + lookup names + +let rec local_binders_length = function + (* Assume that no `{ ... } contexts occur *) + | [] -> 0 + | Topconstr.LocalRawDef _::bl -> 1 + local_binders_length bl + | Topconstr.LocalRawAssum (idl,_,_)::bl -> List.length idl + local_binders_length bl + +let prepare_body (name,annot,args,types,body) rt = + let n = local_binders_length args in +(* Pp.msgnl (str "nb lambda to chop : " ++ str (string_of_int n) ++ fnl () ++Printer.pr_rawconstr rt); *) + let fun_args,rt' = chop_rlambda_n n rt in + (fun_args,rt') + + +let derive_inversion fix_names = + try + (* we first transform the fix_names identifier into their corresponding constant *) + let fix_names_as_constant = + List.map (fun id -> destConst (Tacinterp.constr_of_id (Global.env ()) id)) fix_names + in + (* + Then we check that the graphs have been defined + If one of the graphs haven't been defined + we do nothing + *) + List.iter (fun c -> ignore (find_Function_infos c)) fix_names_as_constant ; + try + Invfun.derive_correctness + Functional_principles_types.make_scheme + functional_induction + fix_names_as_constant + (*i The next call to mk_rel_id is valid since we have just construct the graph + Ensures by : register_built + i*) + (List.map + (fun id -> destInd (Tacinterp.constr_of_id (Global.env ()) (mk_rel_id id))) + fix_names + ) + with e -> + msg_warning + (str "Cannot built inversion information" ++ + if do_observe () then Cerrors.explain_exn e else mt ()) + with _ -> () + +let warning_error names e = + let e_explain e = + match e with + | ToShow e -> spc () ++ Cerrors.explain_exn e + | _ -> if do_observe () then (spc () ++ Cerrors.explain_exn e) else mt () + in + match e with + | Building_graph e -> + Pp.msg_warning + (str "Cannot define graph(s) for " ++ + h 1 (prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) ++ + e_explain e) + | Defining_principle e -> + Pp.msg_warning + (str "Cannot define principle(s) for "++ + h 1 (prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) ++ + e_explain e) + | _ -> anomaly "" + +let error_error names e = + let e_explain e = + match e with + | ToShow e -> spc () ++ Cerrors.explain_exn e + | _ -> if do_observe () then (spc () ++ Cerrors.explain_exn e) else mt () + in + match e with + | Building_graph e -> + errorlabstrm "" + (str "Cannot define graph(s) for " ++ + h 1 (prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) ++ + e_explain e) + | _ -> anomaly "" + +let generate_principle on_error + is_general do_built fix_rec_l recdefs interactive_proof + (continue_proof : int -> Names.constant array -> Term.constr array -> int -> + Tacmach.tactic) : unit = + let names = List.map (function ((_, name),_,_,_,_) -> name) fix_rec_l in + let fun_bodies = List.map2 prepare_body fix_rec_l recdefs in + let funs_args = List.map fst fun_bodies in + let funs_types = List.map (function (_,_,_,types,_) -> types) fix_rec_l in + try + (* We then register the Inductive graphs of the functions *) + Rawterm_to_relation.build_inductive names funs_args funs_types recdefs; + if do_built + then + begin + (*i The next call to mk_rel_id is valid since we have just construct the graph + Ensures by : do_built + i*) + let f_R_mut = Ident (dummy_loc,mk_rel_id (List.nth names 0)) in + let ind_kn = + fst (locate_with_msg + (pr_reference f_R_mut++str ": Not an inductive type!") + locate_ind + f_R_mut) + in + let fname_kn (fname,_,_,_,_) = + let f_ref = Ident fname in + locate_with_msg + (pr_reference f_ref++str ": Not an inductive type!") + locate_constant + f_ref + in + let funs_kn = Array.of_list (List.map fname_kn fix_rec_l) in + let _ = + list_map_i + (fun i x -> + let princ = destConst (Indrec.lookup_eliminator (ind_kn,i) (InProp)) in + let princ_type = Typeops.type_of_constant (Global.env()) princ + in + Functional_principles_types.generate_functional_principle + interactive_proof + princ_type + None + None + funs_kn + i + (continue_proof 0 [|funs_kn.(i)|]) + ) + 0 + fix_rec_l + in + Array.iter (add_Function is_general) funs_kn; + () + end + with e -> + on_error names e + +let register_struct is_rec fixpoint_exprl = + match fixpoint_exprl with + | [((_,fname),_,bl,ret_type,body),_] when not is_rec -> + let ce,imps = + Command.interp_definition + (Flags.boxed_definitions ()) bl None body (Some ret_type) + in + Command.declare_definition + fname (Decl_kinds.Global,Decl_kinds.Definition) + ce imps (fun _ _ -> ()) + | _ -> + let fixpoint_exprl = + List.map (fun ((name,annot,bl,types,body),ntn) -> + ((name,annot,bl,types,Some body),ntn)) fixpoint_exprl in + Command.do_fixpoint fixpoint_exprl (Flags.boxed_definitions()) + +let generate_correction_proof_wf f_ref tcc_lemma_ref + is_mes functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation + (_: int) (_:Names.constant array) (_:Term.constr array) (_:int) : Tacmach.tactic = + Functional_principles_proofs.prove_principle_for_gen + (f_ref,functional_ref,eq_ref) + tcc_lemma_ref is_mes rec_arg_num rec_arg_type relation + + +let register_wf ?(is_mes=false) fname rec_impls wf_rel_expr wf_arg using_lemmas args ret_type body + pre_hook + = + let type_of_f = Topconstr.prod_constr_expr ret_type args in + let rec_arg_num = + let names = + List.map + snd + (Topconstr.names_of_local_assums args) + in + match wf_arg with + | None -> + if List.length names = 1 then 1 + else error "Recursive argument must be specified" + | Some wf_arg -> + list_index (Name wf_arg) names + in + let unbounded_eq = + let f_app_args = + Topconstr.CAppExpl + (dummy_loc, + (None,(Ident (dummy_loc,fname))) , + (List.map + (function + | _,Anonymous -> assert false + | _,Name e -> (Topconstr.mkIdentC e) + ) + (Topconstr.names_of_local_assums args) + ) + ) + in + Topconstr.CApp (dummy_loc,(None,Topconstr.mkRefC (Qualid (dummy_loc,(qualid_of_string "Logic.eq")))), + [(f_app_args,None);(body,None)]) + in + let eq = Topconstr.prod_constr_expr unbounded_eq args in + let hook f_ref tcc_lemma_ref functional_ref eq_ref rec_arg_num rec_arg_type + nb_args relation = + try + pre_hook + (generate_correction_proof_wf f_ref tcc_lemma_ref is_mes + functional_ref eq_ref rec_arg_num rec_arg_type nb_args relation + ); + derive_inversion [fname] + with e -> + (* No proof done *) + () + in + Recdef.recursive_definition + is_mes fname rec_impls + type_of_f + wf_rel_expr + rec_arg_num + eq + hook + using_lemmas + + +let register_mes fname rec_impls wf_mes_expr wf_arg using_lemmas args ret_type body = + let wf_arg_type,wf_arg = + match wf_arg with + | None -> + begin + match args with + | [Topconstr.LocalRawAssum ([(_,Name x)],k,t)] -> t,x + | _ -> error "Recursive argument must be specified" + end + | Some wf_args -> + try + match + List.find + (function + | Topconstr.LocalRawAssum(l,k,t) -> + List.exists + (function (_,Name id) -> id = wf_args | _ -> false) + l + | _ -> false + ) + args + with + | Topconstr.LocalRawAssum(_,k,t) -> t,wf_args + | _ -> assert false + with Not_found -> assert false + in + let ltof = + let make_dir l = make_dirpath (List.map id_of_string (List.rev l)) in + Libnames.Qualid (dummy_loc,Libnames.qualid_of_path + (Libnames.make_path (make_dir ["Arith";"Wf_nat"]) (id_of_string "ltof"))) + in + let fun_from_mes = + let applied_mes = + Topconstr.mkAppC(wf_mes_expr,[Topconstr.mkIdentC wf_arg]) in + Topconstr.mkLambdaC ([(dummy_loc,Name wf_arg)],Topconstr.default_binder_kind,wf_arg_type,applied_mes) + in + let wf_rel_from_mes = + Topconstr.mkAppC(Topconstr.mkRefC ltof,[wf_arg_type;fun_from_mes]) + in + register_wf ~is_mes:true fname rec_impls wf_rel_from_mes (Some wf_arg) + using_lemmas args ret_type body + + +let do_generate_principle on_error register_built interactive_proof fixpoint_exprl = + let recdefs,rec_impls = build_newrecursive fixpoint_exprl in + let _is_struct = + match fixpoint_exprl with + | [(((_,name),Some (Wf (wf_rel,wf_x,using_lemmas)),args,types,body))] -> + let pre_hook = + generate_principle + on_error + true + register_built + fixpoint_exprl + recdefs + true + in + if register_built + then register_wf name rec_impls wf_rel wf_x using_lemmas args types body pre_hook; + false + | [(((_,name),Some (Mes (wf_mes,wf_x,using_lemmas)),args,types,body))] -> + let pre_hook = + generate_principle + on_error + true + register_built + fixpoint_exprl + recdefs + true + in + if register_built + then register_mes name rec_impls wf_mes wf_x using_lemmas args types body pre_hook; + true + | _ -> + let fix_names = + List.map (function ((_,name),_,_,_,_) -> name) fixpoint_exprl + in + let is_one_rec = is_rec fix_names in + let old_fixpoint_exprl = + List.map + (function + | (name,Some (Struct id),args,types,body),_ -> + let annot = + try Some (dummy_loc, id), Topconstr.CStructRec + with Not_found -> + raise (UserError("",str "Cannot find argument " ++ + Ppconstr.pr_id id)) + in + (name,annot,args,types,body),([]:Vernacexpr.decl_notation list) + | (name,None,args,types,body),recdef -> + let names = (Topconstr.names_of_local_assums args) in + if is_one_rec recdef && List.length names > 1 then + user_err_loc + (dummy_loc,"Function", + Pp.str "the recursive argument needs to be specified in Function") + else + let loc, na = List.hd names in + (name,(Some (loc, Nameops.out_name na), Topconstr.CStructRec),args,types,body), + ([]:Vernacexpr.decl_notation list) + | (_,Some (Wf _),_,_,_),_ | (_,Some (Mes _),_,_,_),_-> + error + ("Cannot use mutual definition with well-founded recursion or measure") + ) + (List.combine fixpoint_exprl recdefs) + in + (* ok all the expressions are structural *) + let fix_names = + List.map (function ((_,name),_,_,_,_) -> name) fixpoint_exprl + in + let is_rec = List.exists (is_rec fix_names) recdefs in + if register_built then register_struct is_rec old_fixpoint_exprl; + generate_principle + on_error + false + register_built + fixpoint_exprl + recdefs + interactive_proof + (Functional_principles_proofs.prove_princ_for_struct interactive_proof); + if register_built then derive_inversion fix_names; + true; + in + () + +open Topconstr +let rec add_args id new_args b = + match b with + | CRef r -> + begin match r with + | Libnames.Ident(loc,fname) when fname = id -> + CAppExpl(dummy_loc,(None,r),new_args) + | _ -> b + end + | CFix _ | CCoFix _ -> anomaly "add_args : todo" + | CArrow(loc,b1,b2) -> + CArrow(loc,add_args id new_args b1, add_args id new_args b2) + | CProdN(loc,nal,b1) -> + CProdN(loc, + List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal, + add_args id new_args b1) + | CLambdaN(loc,nal,b1) -> + CLambdaN(loc, + List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal, + add_args id new_args b1) + | CLetIn(loc,na,b1,b2) -> + CLetIn(loc,na,add_args id new_args b1,add_args id new_args b2) + | CAppExpl(loc,(pf,r),exprl) -> + begin + match r with + | Libnames.Ident(loc,fname) when fname = id -> + CAppExpl(loc,(pf,r),new_args@(List.map (add_args id new_args) exprl)) + | _ -> CAppExpl(loc,(pf,r),List.map (add_args id new_args) exprl) + end + | CApp(loc,(pf,b),bl) -> + CApp(loc,(pf,add_args id new_args b), + List.map (fun (e,o) -> add_args id new_args e,o) bl) + | CCases(loc,sty,b_option,cel,cal) -> + CCases(loc,sty,Option.map (add_args id new_args) b_option, + List.map (fun (b,(na,b_option)) -> + add_args id new_args b, + (na,Option.map (add_args id new_args) b_option)) cel, + List.map (fun (loc,cpl,e) -> (loc,cpl,add_args id new_args e)) cal + ) + | CLetTuple(loc,nal,(na,b_option),b1,b2) -> + CLetTuple(loc,nal,(na,Option.map (add_args id new_args) b_option), + add_args id new_args b1, + add_args id new_args b2 + ) + + | CIf(loc,b1,(na,b_option),b2,b3) -> + CIf(loc,add_args id new_args b1, + (na,Option.map (add_args id new_args) b_option), + add_args id new_args b2, + add_args id new_args b3 + ) + | CHole _ -> b + | CPatVar _ -> b + | CEvar _ -> b + | CSort _ -> b + | CCast(loc,b1,CastConv(ck,b2)) -> + CCast(loc,add_args id new_args b1,CastConv(ck,add_args id new_args b2)) + | CCast(loc,b1,CastCoerce) -> + CCast(loc,add_args id new_args b1,CastCoerce) + | CRecord (loc, w, pars) -> + CRecord (loc, + (match w with Some w -> Some (add_args id new_args w) | _ -> None), + List.map (fun (e,o) -> e, add_args id new_args o) pars) + | CNotation _ -> anomaly "add_args : CNotation" + | CGeneralization _ -> anomaly "add_args : CGeneralization" + | CPrim _ -> b + | CDelimiters _ -> anomaly "add_args : CDelimiters" + | CDynamic _ -> anomaly "add_args : CDynamic" +exception Stop of Topconstr.constr_expr + + +(* [chop_n_arrow n t] chops the [n] first arrows in [t] + Acts on Topconstr.constr_expr +*) +let rec chop_n_arrow n t = + if n <= 0 + then t (* If we have already removed all the arrows then return the type *) + else (* If not we check the form of [t] *) + match t with + | Topconstr.CArrow(_,_,t) -> (* If we have an arrow, we discard it and recall [chop_n_arrow] *) + chop_n_arrow (n-1) t + | Topconstr.CProdN(_,nal_ta',t') -> (* If we have a forall, to result are possible : + either we need to discard more than the number of arrows contained + in this product declaration then we just recall [chop_n_arrow] on + the remaining number of arrow to chop and [t'] we discard it and + recall [chop_n_arrow], either this product contains more arrows + than the number we need to chop and then we return the new type + *) + begin + try + let new_n = + let rec aux (n:int) = function + [] -> n + | (nal,k,t'')::nal_ta' -> + let nal_l = List.length nal in + if n >= nal_l + then + aux (n - nal_l) nal_ta' + else + let new_t' = + Topconstr.CProdN(dummy_loc, + ((snd (list_chop n nal)),k,t'')::nal_ta',t') + in + raise (Stop new_t') + in + aux n nal_ta' + in + chop_n_arrow new_n t' + with Stop t -> t + end + | _ -> anomaly "Not enough products" + + +let rec get_args b t : Topconstr.local_binder list * + Topconstr.constr_expr * Topconstr.constr_expr = + match b with + | Topconstr.CLambdaN (loc, (nal_ta), b') -> + begin + let n = + (List.fold_left (fun n (nal,_,_) -> + n+List.length nal) 0 nal_ta ) + in + let nal_tas,b'',t'' = get_args b' (chop_n_arrow n t) in + (List.map (fun (nal,k,ta) -> + (Topconstr.LocalRawAssum (nal,k,ta))) nal_ta)@nal_tas, b'',t'' + end + | _ -> [],b,t + + +let make_graph (f_ref:global_reference) = + let c,c_body = + match f_ref with + | ConstRef c -> + begin try c,Global.lookup_constant c + with Not_found -> + raise (UserError ("",str "Cannot find " ++ Printer.pr_lconstr (mkConst c)) ) + end + | _ -> raise (UserError ("", str "Not a function reference") ) + + in + Dumpglob.pause (); + (match c_body.const_body with + | None -> error "Cannot build a graph over an axiom !" + | Some b -> + let env = Global.env () in + let body = (force b) in + let extern_body,extern_type = + with_full_print + (fun () -> + (Constrextern.extern_constr false env body, + Constrextern.extern_type false env + (Typeops.type_of_constant_type env c_body.const_type) + ) + ) + () + in + let (nal_tas,b,t) = get_args extern_body extern_type in + let expr_list = + match b with + | Topconstr.CFix(loc,l_id,fixexprl) -> + let l = + List.map + (fun (id,(n,recexp),bl,t,b) -> + let loc, rec_id = Option.get n in + let new_args = + List.flatten + (List.map + (function + | Topconstr.LocalRawDef (na,_)-> [] + | Topconstr.LocalRawAssum (nal,_,_) -> + List.map + (fun (loc,n) -> + CRef(Libnames.Ident(loc, Nameops.out_name n))) + nal + ) + nal_tas + ) + in + let b' = add_args (snd id) new_args b in + (id, Some (Struct rec_id),nal_tas@bl,t,b') + ) + fixexprl + in + l + | _ -> + let id = id_of_label (con_label c) in + [((dummy_loc,id),None,nal_tas,t,b)] + in + do_generate_principle error_error false false expr_list; + (* We register the infos *) + let mp,dp,_ = repr_con c in + List.iter + (fun ((_,id),_,_,_,_) -> add_Function false (make_con mp dp (label_of_id id))) + expr_list); + Dumpglob.continue () + + +(* let make_graph _ = assert false *) + +let do_generate_principle = do_generate_principle warning_error true + + diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml new file mode 100644 index 00000000..0f048f59 --- /dev/null +++ b/plugins/funind/indfun_common.ml @@ -0,0 +1,558 @@ +open Names +open Pp + +open Libnames + +let mk_prefix pre id = id_of_string (pre^(string_of_id id)) +let mk_rel_id = mk_prefix "R_" +let mk_correct_id id = Nameops.add_suffix (mk_rel_id id) "_correct" +let mk_complete_id id = Nameops.add_suffix (mk_rel_id id) "_complete" +let mk_equation_id id = Nameops.add_suffix id "_equation" + +let msgnl m = + () + +let invalid_argument s = raise (Invalid_argument s) + + +let fresh_id avoid s = Namegen.next_ident_away_in_goal (id_of_string s) avoid + +let fresh_name avoid s = Name (fresh_id avoid s) + +let get_name avoid ?(default="H") = function + | Anonymous -> fresh_name avoid default + | Name n -> Name n + +let array_get_start a = + try + Array.init + (Array.length a - 1) + (fun i -> a.(i)) + with Invalid_argument "index out of bounds" -> + invalid_argument "array_get_start" + +let id_of_name = function + Name id -> id + | _ -> raise Not_found + +let locate ref = + let (loc,qid) = qualid_of_reference ref in + Nametab.locate qid + +let locate_ind ref = + match locate ref with + | IndRef x -> x + | _ -> raise Not_found + +let locate_constant ref = + match locate ref with + | ConstRef x -> x + | _ -> raise Not_found + + +let locate_with_msg msg f x = + try + f x + with + | Not_found -> raise (Util.UserError("", msg)) + | e -> raise e + + +let filter_map filter f = + let rec it = function + | [] -> [] + | e::l -> + if filter e + then + (f e) :: it l + else it l + in + it + + +let chop_rlambda_n = + let rec chop_lambda_n acc n rt = + if n == 0 + then List.rev acc,rt + else + match rt with + | Rawterm.RLambda(_,name,k,t,b) -> chop_lambda_n ((name,t,false)::acc) (n-1) b + | Rawterm.RLetIn(_,name,v,b) -> chop_lambda_n ((name,v,true)::acc) (n-1) b + | _ -> + raise (Util.UserError("chop_rlambda_n", + str "chop_rlambda_n: Not enough Lambdas")) + in + chop_lambda_n [] + +let chop_rprod_n = + let rec chop_prod_n acc n rt = + if n == 0 + then List.rev acc,rt + else + match rt with + | Rawterm.RProd(_,name,k,t,b) -> chop_prod_n ((name,t)::acc) (n-1) b + | _ -> raise (Util.UserError("chop_rprod_n",str "chop_rprod_n: Not enough products")) + in + chop_prod_n [] + + + +let list_union_eq eq_fun l1 l2 = + let rec urec = function + | [] -> l2 + | a::l -> if List.exists (eq_fun a) l2 then urec l else a::urec l + in + urec l1 + +let list_add_set_eq eq_fun x l = + if List.exists (eq_fun x) l then l else x::l + + + + +let const_of_id id = + let _,princ_ref = + qualid_of_reference (Libnames.Ident (Util.dummy_loc,id)) + in + try Nametab.locate_constant princ_ref + with Not_found -> Util.error ("cannot find "^ string_of_id id) + +let def_of_const t = + match (Term.kind_of_term t) with + Term.Const sp -> + (try (match (Global.lookup_constant sp) with + {Declarations.const_body=Some c} -> Declarations.force c + |_ -> assert false) + with _ -> assert false) + |_ -> assert false + +let coq_constant s = + Coqlib.gen_constant_in_modules "RecursiveDefinition" + Coqlib.init_modules s;; + +let constant sl s = + constr_of_global + (Nametab.locate (make_qualid(Names.make_dirpath + (List.map id_of_string (List.rev sl))) + (id_of_string s)));; + +let find_reference sl s = + (Nametab.locate (make_qualid(Names.make_dirpath + (List.map id_of_string (List.rev sl))) + (id_of_string s)));; + +let eq = lazy(coq_constant "eq") +let refl_equal = lazy(coq_constant "eq_refl") + +(*****************************************************************) +(* Copy of the standart save mechanism but without the much too *) +(* slow reduction function *) +(*****************************************************************) +open Declarations +open Entries +open Decl_kinds +open Declare +let definition_message id = + Flags.if_verbose message ((string_of_id id) ^ " is defined") + + +let save with_clean id const (locality,kind) hook = + let {const_entry_body = pft; + const_entry_type = tpo; + const_entry_opaque = opacity } = const in + let l,r = match locality with + | Local when Lib.sections_are_opened () -> + let k = logical_kind_of_goal_kind kind in + let c = SectionLocalDef (pft, tpo, opacity) in + let _ = declare_variable id (Lib.cwd(), c, k) in + (Local, VarRef id) + | Local -> + let k = logical_kind_of_goal_kind kind in + let kn = declare_constant id (DefinitionEntry const, k) in + (Global, ConstRef kn) + | Global -> + let k = logical_kind_of_goal_kind kind in + let kn = declare_constant id (DefinitionEntry const, k) in + (Global, ConstRef kn) in + if with_clean then Pfedit.delete_current_proof (); + hook l r; + definition_message id + + + + +let extract_pftreestate pts = + let pfterm,subgoals = Refiner.extract_open_pftreestate pts in + let tpfsigma = Refiner.evc_of_pftreestate pts in + let exl = Evarutil.non_instantiated tpfsigma in + if subgoals <> [] or exl <> [] then + Util.errorlabstrm "extract_proof" + (if subgoals <> [] then + str "Attempt to save an incomplete proof" + else + str "Attempt to save a proof with existential variables still non-instantiated"); + let env = Global.env_of_context (Refiner.proof_of_pftreestate pts).Proof_type.goal.Evd.evar_hyps in + env,tpfsigma,pfterm + + +let nf_betaiotazeta = + let clos_norm_flags flgs env sigma t = + Closure.norm_val (Closure.create_clos_infos flgs env) (Closure.inject (Reductionops.nf_evar sigma t)) in + clos_norm_flags Closure.betaiotazeta + +let nf_betaiota = + let clos_norm_flags flgs env sigma t = + Closure.norm_val (Closure.create_clos_infos flgs env) (Closure.inject (Reductionops.nf_evar sigma t)) in + clos_norm_flags Closure.betaiota + +let cook_proof do_reduce = + let pfs = Pfedit.get_pftreestate () +(* and ident = Pfedit.get_current_proof_name () *) + and (ident,strength,concl,hook) = Pfedit.current_proof_statement () in + let env,sigma,pfterm = extract_pftreestate pfs in + let pfterm = + if do_reduce + then nf_betaiota env sigma pfterm + else pfterm + in + (ident, + ({ const_entry_body = pfterm; + const_entry_type = Some concl; + const_entry_opaque = false; + const_entry_boxed = false}, + strength, hook)) + + +let new_save_named opacity = + let id,(const,persistence,hook) = cook_proof true in + let const = { const with const_entry_opaque = opacity } in + save true id const persistence hook + +let get_proof_clean do_reduce = + let result = cook_proof do_reduce in + Pfedit.delete_current_proof (); + result + +let with_full_print f a = + let old_implicit_args = Impargs.is_implicit_args () + and old_strict_implicit_args = Impargs.is_strict_implicit_args () + and old_contextual_implicit_args = Impargs.is_contextual_implicit_args () in + let old_rawprint = !Flags.raw_print in + Flags.raw_print := true; + Impargs.make_implicit_args false; + Impargs.make_strict_implicit_args false; + Impargs.make_contextual_implicit_args false; + Impargs.make_contextual_implicit_args false; + Dumpglob.pause (); + try + let res = f a in + Impargs.make_implicit_args old_implicit_args; + Impargs.make_strict_implicit_args old_strict_implicit_args; + Impargs.make_contextual_implicit_args old_contextual_implicit_args; + Flags.raw_print := old_rawprint; + Dumpglob.continue (); + res + with + | e -> + Impargs.make_implicit_args old_implicit_args; + Impargs.make_strict_implicit_args old_strict_implicit_args; + Impargs.make_contextual_implicit_args old_contextual_implicit_args; + Flags.raw_print := old_rawprint; + Dumpglob.continue (); + raise e + + + + + + +(**********************) + +type function_info = + { + function_constant : constant; + graph_ind : inductive; + equation_lemma : constant option; + correctness_lemma : constant option; + completeness_lemma : constant option; + rect_lemma : constant option; + rec_lemma : constant option; + prop_lemma : constant option; + is_general : bool; (* Has this function been defined using general recursive definition *) + } + + +(* type function_db = function_info list *) + +(* let function_table = ref ([] : function_db) *) + + +let from_function = ref Cmap.empty +let from_graph = ref Indmap.empty +(* +let rec do_cache_info finfo = function + | [] -> raise Not_found + | (finfo'::finfos as l) -> + if finfo' == finfo then l + else if finfo'.function_constant = finfo.function_constant + then finfo::finfos + else + let res = do_cache_info finfo finfos in + if res == finfos then l else finfo'::l + + +let cache_Function (_,(finfos)) = + let new_tbl = + try do_cache_info finfos !function_table + with Not_found -> finfos::!function_table + in + if new_tbl != !function_table + then function_table := new_tbl +*) + +let cache_Function (_,finfos) = + from_function := Cmap.add finfos.function_constant finfos !from_function; + from_graph := Indmap.add finfos.graph_ind finfos !from_graph + + +let load_Function _ = cache_Function +let open_Function _ = cache_Function +let subst_Function (subst,finfos) = + let do_subst_con c = fst (Mod_subst.subst_con subst c) + and do_subst_ind (kn,i) = (Mod_subst.subst_ind subst kn,i) + in + let function_constant' = do_subst_con finfos.function_constant in + let graph_ind' = do_subst_ind finfos.graph_ind in + let equation_lemma' = Option.smartmap do_subst_con finfos.equation_lemma in + let correctness_lemma' = Option.smartmap do_subst_con finfos.correctness_lemma in + let completeness_lemma' = Option.smartmap do_subst_con finfos.completeness_lemma in + let rect_lemma' = Option.smartmap do_subst_con finfos.rect_lemma in + let rec_lemma' = Option.smartmap do_subst_con finfos.rec_lemma in + let prop_lemma' = Option.smartmap do_subst_con finfos.prop_lemma in + if function_constant' == finfos.function_constant && + graph_ind' == finfos.graph_ind && + equation_lemma' == finfos.equation_lemma && + correctness_lemma' == finfos.correctness_lemma && + completeness_lemma' == finfos.completeness_lemma && + rect_lemma' == finfos.rect_lemma && + rec_lemma' == finfos.rec_lemma && + prop_lemma' == finfos.prop_lemma + then finfos + else + { function_constant = function_constant'; + graph_ind = graph_ind'; + equation_lemma = equation_lemma' ; + correctness_lemma = correctness_lemma' ; + completeness_lemma = completeness_lemma' ; + rect_lemma = rect_lemma' ; + rec_lemma = rec_lemma'; + prop_lemma = prop_lemma'; + is_general = finfos.is_general + } + +let classify_Function infos = Libobject.Substitute infos + + +let discharge_Function (_,finfos) = + let function_constant' = Lib.discharge_con finfos.function_constant + and graph_ind' = Lib.discharge_inductive finfos.graph_ind + and equation_lemma' = Option.smartmap Lib.discharge_con finfos.equation_lemma + and correctness_lemma' = Option.smartmap Lib.discharge_con finfos.correctness_lemma + and completeness_lemma' = Option.smartmap Lib.discharge_con finfos.completeness_lemma + and rect_lemma' = Option.smartmap Lib.discharge_con finfos.rect_lemma + and rec_lemma' = Option.smartmap Lib.discharge_con finfos.rec_lemma + and prop_lemma' = Option.smartmap Lib.discharge_con finfos.prop_lemma + in + if function_constant' == finfos.function_constant && + graph_ind' == finfos.graph_ind && + equation_lemma' == finfos.equation_lemma && + correctness_lemma' == finfos.correctness_lemma && + completeness_lemma' == finfos.completeness_lemma && + rect_lemma' == finfos.rect_lemma && + rec_lemma' == finfos.rec_lemma && + prop_lemma' == finfos.prop_lemma + then Some finfos + else + Some { function_constant = function_constant' ; + graph_ind = graph_ind' ; + equation_lemma = equation_lemma' ; + correctness_lemma = correctness_lemma' ; + completeness_lemma = completeness_lemma'; + rect_lemma = rect_lemma'; + rec_lemma = rec_lemma'; + prop_lemma = prop_lemma' ; + is_general = finfos.is_general + } + +open Term +let pr_info f_info = + str "function_constant := " ++ Printer.pr_lconstr (mkConst f_info.function_constant)++ fnl () ++ + str "function_constant_type := " ++ + (try Printer.pr_lconstr (Global.type_of_global (ConstRef f_info.function_constant)) with _ -> mt ()) ++ fnl () ++ + str "equation_lemma := " ++ (Option.fold_right (fun v acc -> Printer.pr_lconstr (mkConst v)) f_info.equation_lemma (mt ()) ) ++ fnl () ++ + str "completeness_lemma :=" ++ (Option.fold_right (fun v acc -> Printer.pr_lconstr (mkConst v)) f_info.completeness_lemma (mt ()) ) ++ fnl () ++ + str "correctness_lemma := " ++ (Option.fold_right (fun v acc -> Printer.pr_lconstr (mkConst v)) f_info.correctness_lemma (mt ()) ) ++ fnl () ++ + str "rect_lemma := " ++ (Option.fold_right (fun v acc -> Printer.pr_lconstr (mkConst v)) f_info.rect_lemma (mt ()) ) ++ fnl () ++ + str "rec_lemma := " ++ (Option.fold_right (fun v acc -> Printer.pr_lconstr (mkConst v)) f_info.rec_lemma (mt ()) ) ++ fnl () ++ + str "prop_lemma := " ++ (Option.fold_right (fun v acc -> Printer.pr_lconstr (mkConst v)) f_info.prop_lemma (mt ()) ) ++ fnl () ++ + str "graph_ind := " ++ Printer.pr_lconstr (mkInd f_info.graph_ind) ++ fnl () + +let pr_table tb = + let l = Cmap.fold (fun k v acc -> v::acc) tb [] in + Util.prlist_with_sep fnl pr_info l + +let in_Function,out_Function = + Libobject.declare_object + {(Libobject.default_object "FUNCTIONS_DB") with + Libobject.cache_function = cache_Function; + Libobject.load_function = load_Function; + Libobject.classify_function = classify_Function; + Libobject.subst_function = subst_Function; + Libobject.discharge_function = discharge_Function +(* Libobject.open_function = open_Function; *) + } + + + +(* Synchronisation with reset *) +let freeze () = + !from_function,!from_graph +let unfreeze (functions,graphs) = +(* Pp.msgnl (str "unfreezing function_table : " ++ pr_table l); *) + from_function := functions; + from_graph := graphs + +let init () = +(* Pp.msgnl (str "reseting function_table"); *) + from_function := Cmap.empty; + from_graph := Indmap.empty + +let _ = + Summary.declare_summary "functions_db_sum" + { Summary.freeze_function = freeze; + Summary.unfreeze_function = unfreeze; + Summary.init_function = init } + +let find_or_none id = + try Some + (match Nametab.locate (qualid_of_ident id) with ConstRef c -> c | _ -> Util.anomaly "Not a constant" + ) + with Not_found -> None + + + +let find_Function_infos f = + Cmap.find f !from_function + + +let find_Function_of_graph ind = + Indmap.find ind !from_graph + +let update_Function finfo = +(* Pp.msgnl (pr_info finfo); *) + Lib.add_anonymous_leaf (in_Function finfo) + + +let add_Function is_general f = + let f_id = id_of_label (con_label f) in + let equation_lemma = find_or_none (mk_equation_id f_id) + and correctness_lemma = find_or_none (mk_correct_id f_id) + and completeness_lemma = find_or_none (mk_complete_id f_id) + and rect_lemma = find_or_none (Nameops.add_suffix f_id "_rect") + and rec_lemma = find_or_none (Nameops.add_suffix f_id "_rec") + and prop_lemma = find_or_none (Nameops.add_suffix f_id "_ind") + and graph_ind = + match Nametab.locate (qualid_of_ident (mk_rel_id f_id)) + with | IndRef ind -> ind | _ -> Util.anomaly "Not an inductive" + in + let finfos = + { function_constant = f; + equation_lemma = equation_lemma; + completeness_lemma = completeness_lemma; + correctness_lemma = correctness_lemma; + rect_lemma = rect_lemma; + rec_lemma = rec_lemma; + prop_lemma = prop_lemma; + graph_ind = graph_ind; + is_general = is_general + + } + in + update_Function finfos + +let pr_table () = pr_table !from_function +(*********************************) +(* Debuging *) +let functional_induction_rewrite_dependent_proofs = ref true +let function_debug = ref false +open Goptions + +let functional_induction_rewrite_dependent_proofs_sig = + { + optsync = false; + optname = "Functional Induction Rewrite Dependent"; + optkey = ["Functional";"Induction";"Rewrite";"Dependent"]; + optread = (fun () -> !functional_induction_rewrite_dependent_proofs); + optwrite = (fun b -> functional_induction_rewrite_dependent_proofs := b) + } +let _ = declare_bool_option functional_induction_rewrite_dependent_proofs_sig + +let do_rewrite_dependent () = !functional_induction_rewrite_dependent_proofs = true + +let function_debug_sig = + { + optsync = false; + optname = "Function debug"; + optkey = ["Function_debug"]; + optread = (fun () -> !function_debug); + optwrite = (fun b -> function_debug := b) + } + +let _ = declare_bool_option function_debug_sig + + +let do_observe () = + !function_debug = true + + + +let strict_tcc = ref false +let is_strict_tcc () = !strict_tcc +let strict_tcc_sig = + { + optsync = false; + optname = "Raw Function Tcc"; + optkey = ["Function_raw_tcc"]; + optread = (fun () -> !strict_tcc); + optwrite = (fun b -> strict_tcc := b) + } + +let _ = declare_bool_option strict_tcc_sig + + +exception Building_graph of exn +exception Defining_principle of exn +exception ToShow of exn + +let init_constant dir s = + try + Coqlib.gen_constant "Function" dir s + with e -> raise (ToShow e) + +let jmeq () = + try + (Coqlib.check_required_library ["Coq";"Logic";"JMeq"]; + init_constant ["Logic";"JMeq"] "JMeq") + with e -> raise (ToShow e) + +let jmeq_rec () = + try + Coqlib.check_required_library ["Coq";"Logic";"JMeq"]; + init_constant ["Logic";"JMeq"] "JMeq_rec" + with e -> raise (ToShow e) + +let jmeq_refl () = + try + Coqlib.check_required_library ["Coq";"Logic";"JMeq"]; + init_constant ["Logic";"JMeq"] "JMeq_refl" + with e -> raise (ToShow e) diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli new file mode 100644 index 00000000..6f6607fc --- /dev/null +++ b/plugins/funind/indfun_common.mli @@ -0,0 +1,121 @@ +open Names +open Pp + +(* + The mk_?_id function build different name w.r.t. a function + Each of their use is justified in the code +*) +val mk_rel_id : identifier -> identifier +val mk_correct_id : identifier -> identifier +val mk_complete_id : identifier -> identifier +val mk_equation_id : identifier -> identifier + + +val msgnl : std_ppcmds -> unit + +val invalid_argument : string -> 'a + +val fresh_id : identifier list -> string -> identifier +val fresh_name : identifier list -> string -> name +val get_name : identifier list -> ?default:string -> name -> name + +val array_get_start : 'a array -> 'a array + +val id_of_name : name -> identifier + +val locate_ind : Libnames.reference -> inductive +val locate_constant : Libnames.reference -> constant +val locate_with_msg : + Pp.std_ppcmds -> (Libnames.reference -> 'a) -> + Libnames.reference -> 'a + +val filter_map : ('a -> bool) -> ('a -> 'b) -> 'a list -> 'b list +val list_union_eq : + ('a -> 'a -> bool) -> 'a list -> 'a list -> 'a list +val list_add_set_eq : + ('a -> 'a -> bool) -> 'a -> 'a list -> 'a list + +val chop_rlambda_n : int -> Rawterm.rawconstr -> + (name*Rawterm.rawconstr*bool) list * Rawterm.rawconstr + +val chop_rprod_n : int -> Rawterm.rawconstr -> + (name*Rawterm.rawconstr) list * Rawterm.rawconstr + +val def_of_const : Term.constr -> Term.constr +val eq : Term.constr Lazy.t +val refl_equal : Term.constr Lazy.t +val const_of_id: identifier -> constant +val jmeq : unit -> Term.constr +val jmeq_refl : unit -> Term.constr + +(* [save_named] is a copy of [Command.save_named] but uses + [nf_betaiotazeta] instead of [nf_betaiotaevar_preserving_vm_cast] + + + + DON'T USE IT if you cannot ensure that there is no VMcast in the proof + +*) + +(* val nf_betaiotazeta : Reductionops.reduction_function *) + +val new_save_named : bool -> unit + +val save : bool -> identifier -> Entries.definition_entry -> Decl_kinds.goal_kind -> + Tacexpr.declaration_hook -> unit + +(* [get_proof_clean do_reduce] : returns the proof name, definition, kind and hook and + abort the proof +*) +val get_proof_clean : bool -> + Names.identifier * + (Entries.definition_entry * Decl_kinds.goal_kind * + Tacexpr.declaration_hook) + + + +(* [with_full_print f a] applies [f] to [a] in full printing environment + + This function preserves the print settings +*) +val with_full_print : ('a -> 'b) -> 'a -> 'b + + +(*****************) + +type function_info = + { + function_constant : constant; + graph_ind : inductive; + equation_lemma : constant option; + correctness_lemma : constant option; + completeness_lemma : constant option; + rect_lemma : constant option; + rec_lemma : constant option; + prop_lemma : constant option; + is_general : bool; + } + +val find_Function_infos : constant -> function_info +val find_Function_of_graph : inductive -> function_info +(* WARNING: To be used just after the graph definition !!! *) +val add_Function : bool -> constant -> unit + +val update_Function : function_info -> unit + + +(** debugging *) +val pr_info : function_info -> Pp.std_ppcmds +val pr_table : unit -> Pp.std_ppcmds + + +(* val function_debug : bool ref *) +val do_observe : unit -> bool +val do_rewrite_dependent : unit -> bool + +(* To localize pb *) +exception Building_graph of exn +exception Defining_principle of exn +exception ToShow of exn + +val is_strict_tcc : unit -> bool diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml new file mode 100644 index 00000000..8c22265d --- /dev/null +++ b/plugins/funind/invfun.ml @@ -0,0 +1,1020 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* hov 1 (Ppconstr.pr_id id ++ str " := " ++ Pp.cut () ++ prc c) + | loc, Rawterm.AnonHyp n, c -> hov 1 (int n ++ str " := " ++ Pp.cut () ++ prc c) + +let pr_bindings prc prlc = function + | Rawterm.ImplicitBindings l -> + brk (1,1) ++ str "with" ++ brk (1,1) ++ + Util.prlist_with_sep spc prc l + | Rawterm.ExplicitBindings l -> + brk (1,1) ++ str "with" ++ brk (1,1) ++ + Util.prlist_with_sep spc (fun b -> str"(" ++ pr_binding prlc b ++ str")") l + | Rawterm.NoBindings -> mt () + + +let pr_with_bindings prc prlc (c,bl) = + prc c ++ hv 0 (pr_bindings prc prlc bl) + + + +let pr_constr_with_binding prc (c,bl) : Pp.std_ppcmds = + pr_with_bindings prc prc (c,bl) + +(* The local debuging mechanism *) +let msgnl = Pp.msgnl + +let observe strm = + if do_observe () + then Pp.msgnl strm + else () + +let observennl strm = + if do_observe () + then begin Pp.msg strm;Pp.pp_flush () end + else () + + +let do_observe_tac s tac g = + let goal = begin try (Printer.pr_goal (sig_it g)) with _ -> assert false end in + try + let v = tac g in msgnl (goal ++ fnl () ++ s ++(str " ")++(str "finished")); v + with e -> + msgnl (str "observation "++ s++str " raised exception " ++ + Cerrors.explain_exn e ++ str " on goal " ++ goal ); + raise e;; + + +let observe_tac s tac g = + if do_observe () + then do_observe_tac (str s) tac g + else tac g + +(* [nf_zeta] $\zeta$-normalization of a term *) +let nf_zeta = + Reductionops.clos_norm_flags (Closure.RedFlags.mkflags [Closure.RedFlags.fZETA]) + Environ.empty_env + Evd.empty + + +(* [id_to_constr id] finds the term associated to [id] in the global environment *) +let id_to_constr id = + try + Tacinterp.constr_of_id (Global.env ()) id + with Not_found -> + raise (UserError ("",str "Cannot find " ++ Ppconstr.pr_id id)) + +(* [generate_type g_to_f f graph i] build the completeness (resp. correctness) lemma type if [g_to_f = true] + (resp. g_to_f = false) where [graph] is the graph of [f] and is the [i]th function in the block. + + [generate_type true f i] returns + \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, + graph\ x_1\ldots x_n\ res \rightarrow res = fv \] decomposed as the context and the conclusion + + [generate_type false f i] returns + \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, + res = fv \rightarrow graph\ x_1\ldots x_n\ res\] decomposed as the context and the conclusion + *) + +let generate_type g_to_f f graph i = + (*i we deduce the number of arguments of the function and its returned type from the graph i*) + let graph_arity = Inductive.type_of_inductive (Global.env()) (Global.lookup_inductive (destInd graph)) in + let ctxt,_ = decompose_prod_assum graph_arity in + let fun_ctxt,res_type = + match ctxt with + | [] | [_] -> anomaly "Not a valid context" + | (_,_,res_type)::fun_ctxt -> fun_ctxt,res_type + in + let nb_args = List.length fun_ctxt in + let args_from_decl i decl = + match decl with + | (_,Some _,_) -> incr i; failwith "args_from_decl" + | _ -> let j = !i in incr i;mkRel (nb_args - j + 1) + in + (*i We need to name the vars [res] and [fv] i*) + let res_id = + Namegen.next_ident_away_in_goal + (id_of_string "res") + (map_succeed (function (Name id,_,_) -> id | (Anonymous,_,_) -> failwith "") fun_ctxt) + in + let fv_id = + Namegen.next_ident_away_in_goal + (id_of_string "fv") + (res_id::(map_succeed (function (Name id,_,_) -> id | (Anonymous,_,_) -> failwith "Anonymous!") fun_ctxt)) + in + (*i we can then type the argument to be applied to the function [f] i*) + let args_as_rels = + let i = ref 0 in + Array.of_list ((map_succeed (args_from_decl i) (List.rev fun_ctxt))) + in + let args_as_rels = Array.map Termops.pop args_as_rels in + (*i + the hypothesis [res = fv] can then be computed + We will need to lift it by one in order to use it as a conclusion + i*) + let res_eq_f_of_args = + mkApp(Coqlib.build_coq_eq (),[|lift 2 res_type;mkRel 1;mkRel 2|]) + in + (*i + The hypothesis [graph\ x_1\ldots x_n\ res] can then be computed + We will need to lift it by one in order to use it as a conclusion + i*) + let graph_applied = + let args_and_res_as_rels = + let i = ref 0 in + Array.of_list ((map_succeed (args_from_decl i) (List.rev ((Name res_id,None,res_type)::fun_ctxt))) ) + in + let args_and_res_as_rels = + Array.mapi (fun i c -> if i <> Array.length args_and_res_as_rels - 1 then lift 1 c else c) args_and_res_as_rels + in + mkApp(graph,args_and_res_as_rels) + in + (*i The [pre_context] is the defined to be the context corresponding to + \[\forall (x_1:t_1)\ldots(x_n:t_n), let fv := f x_1\ldots x_n in, forall res, \] + i*) + let pre_ctxt = + (Name res_id,None,lift 1 res_type)::(Name fv_id,Some (mkApp(mkConst f,args_as_rels)),res_type)::fun_ctxt + in + (*i and we can return the solution depending on which lemma type we are defining i*) + if g_to_f + then (Anonymous,None,graph_applied)::pre_ctxt,(lift 1 res_eq_f_of_args) + else (Anonymous,None,res_eq_f_of_args)::pre_ctxt,(lift 1 graph_applied) + + +(* + [find_induction_principle f] searches and returns the [body] and the [type] of [f_rect] + + WARNING: while convertible, [type_of body] and [type] can be non equal +*) +let find_induction_principle f = + let f_as_constant = match kind_of_term f with + | Const c' -> c' + | _ -> error "Must be used with a function" + in + let infos = find_Function_infos f_as_constant in + match infos.rect_lemma with + | None -> raise Not_found + | Some rect_lemma -> + let rect_lemma = mkConst rect_lemma in + let typ = Typing.type_of (Global.env ()) Evd.empty rect_lemma in + rect_lemma,typ + + + +(* let fname = *) +(* match kind_of_term f with *) +(* | Const c' -> *) +(* id_of_label (con_label c') *) +(* | _ -> error "Must be used with a function" *) +(* in *) + +(* let princ_name = *) +(* ( *) +(* Indrec.make_elimination_ident *) +(* fname *) +(* InType *) +(* ) *) +(* in *) +(* let c = (\* mkConst(mk_from_const (destConst f) princ_name ) in *\) failwith "" in *) +(* c,Typing.type_of (Global.env ()) Evd.empty c *) + + +let rec generate_fresh_id x avoid i = + if i == 0 + then [] + else + let id = Namegen.next_ident_away_in_goal x avoid in + id::(generate_fresh_id x (id::avoid) (pred i)) + + +(* [prove_fun_correct functional_induction funs_constr graphs_constr schemes lemmas_types_infos i ] + is the tactic used to prove correctness lemma. + + [functional_induction] is the tactic defined in [indfun] (dependency problem) + [funs_constr], [graphs_constr] [schemes] [lemmas_types_infos] are the mutually recursive functions + (resp. graphs of the functions and principles and correctness lemma types) to prove correct. + + [i] is the indice of the function to prove correct + + The lemma to prove if suppose to have been generated by [generate_type] (in $\zeta$ normal form that is + it looks like~: + [\forall (x_1:t_1)\ldots(x_n:t_n), forall res, + res = f x_1\ldots x_n in, \rightarrow graph\ x_1\ldots x_n\ res] + + + The sketch of the proof is the following one~: + \begin{enumerate} + \item intros until $x_n$ + \item $functional\ induction\ (f.(i)\ x_1\ldots x_n)$ using schemes.(i) + \item for each generated branch intro [res] and [hres :res = f x_1\ldots x_n], rewrite [hres] and the + apply the corresponding constructor of the corresponding graph inductive. + \end{enumerate} + +*) +let prove_fun_correct functional_induction funs_constr graphs_constr schemes lemmas_types_infos i : tactic = + fun g -> + (* first of all we recreate the lemmas types to be used as predicates of the induction principle + that is~: + \[fun (x_1:t_1)\ldots(x_n:t_n)=> fun fv => fun res => res = fv \rightarrow graph\ x_1\ldots x_n\ res\] + *) + let lemmas = + Array.map + (fun (_,(ctxt,concl)) -> + match ctxt with + | [] | [_] | [_;_] -> anomaly "bad context" + | hres::res::(x,_,t)::ctxt -> + Termops.it_mkLambda_or_LetIn + ~init:(Termops.it_mkProd_or_LetIn ~init:concl [hres;res]) + ((x,None,t)::ctxt) + ) + lemmas_types_infos + in + (* we the get the definition of the graphs block *) + let graph_ind = destInd graphs_constr.(i) in + let kn = fst graph_ind in + let mib,_ = Global.lookup_inductive graph_ind in + (* and the principle to use in this lemma in $\zeta$ normal form *) + let f_principle,princ_type = schemes.(i) in + let princ_type = nf_zeta princ_type in + let princ_infos = Tactics.compute_elim_sig princ_type in + (* The number of args of the function is then easilly computable *) + let nb_fun_args = nb_prod (pf_concl g) - 2 in + let args_names = generate_fresh_id (id_of_string "x") [] nb_fun_args in + let ids = args_names@(pf_ids_of_hyps g) in + (* Since we cannot ensure that the funcitonnal principle is defined in the + environement and due to the bug #1174, we will need to pose the principle + using a name + *) + let principle_id = Namegen.next_ident_away_in_goal (id_of_string "princ") ids in + let ids = principle_id :: ids in + (* We get the branches of the principle *) + let branches = List.rev princ_infos.branches in + (* and built the intro pattern for each of them *) + let intro_pats = + List.map + (fun (_,_,br_type) -> + List.map + (fun id -> dummy_loc, Genarg.IntroIdentifier id) + (generate_fresh_id (id_of_string "y") ids (List.length (fst (decompose_prod_assum br_type)))) + ) + branches + in + (* before building the full intro pattern for the principle *) + let pat = Some (dummy_loc,Genarg.IntroOrAndPattern intro_pats) in + let eq_ind = Coqlib.build_coq_eq () in + let eq_construct = mkConstruct((destInd eq_ind),1) in + (* The next to referencies will be used to find out which constructor to apply in each branch *) + let ind_number = ref 0 + and min_constr_number = ref 0 in + (* The tactic to prove the ith branch of the principle *) + let prove_branche i g = + (* We get the identifiers of this branch *) + let this_branche_ids = + List.fold_right + (fun (_,pat) acc -> + match pat with + | Genarg.IntroIdentifier id -> Idset.add id acc + | _ -> anomaly "Not an identifier" + ) + (List.nth intro_pats (pred i)) + Idset.empty + in + (* and get the real args of the branch by unfolding the defined constant *) + let pre_args,pre_tac = + List.fold_right + (fun (id,b,t) (pre_args,pre_tac) -> + if Idset.mem id this_branche_ids + then + match b with + | None -> (id::pre_args,pre_tac) + | Some b -> + (pre_args, + tclTHEN (h_reduce (Rawterm.Unfold([Rawterm.all_occurrences_expr,EvalVarRef id])) allHyps) pre_tac + ) + + else (pre_args,pre_tac) + ) + (pf_hyps g) + ([],tclIDTAC) + in + (* + We can then recompute the arguments of the constructor. + For each [hid] introduced by this branch, if [hid] has type + $forall res, res=fv -> graph.(j)\ x_1\ x_n res$ the corresponding arguments of the constructor are + [ fv (hid fv (refl_equal fv)) ]. + + If [hid] has another type the corresponding argument of the constructor is [hid] + *) + let constructor_args = + List.fold_right + (fun hid acc -> + let type_of_hid = pf_type_of g (mkVar hid) in + match kind_of_term type_of_hid with + | Prod(_,_,t') -> + begin + match kind_of_term t' with + | Prod(_,t'',t''') -> + begin + match kind_of_term t'',kind_of_term t''' with + | App(eq,args), App(graph',_) + when + (eq_constr eq eq_ind) && + array_exists (eq_constr graph') graphs_constr -> + ((mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);args.(2)|]))|])) + ::args.(2)::acc) + | _ -> mkVar hid :: acc + end + | _ -> mkVar hid :: acc + end + | _ -> mkVar hid :: acc + ) pre_args [] + in + (* in fact we must also add the parameters to the constructor args *) + let constructor_args = + let params_id = fst (list_chop princ_infos.nparams args_names) in + (List.map mkVar params_id)@(List.rev constructor_args) + in + (* We then get the constructor corresponding to this branch and + modifies the references has needed i.e. + if the constructor is the last one of the current inductive then + add one the number of the inductive to take and add the number of constructor of the previous + graph to the minimal constructor number + *) + let constructor = + let constructor_num = i - !min_constr_number in + let length = Array.length (mib.Declarations.mind_packets.(!ind_number).Declarations.mind_consnames) in + if constructor_num <= length + then + begin + (kn,!ind_number),constructor_num + end + else + begin + incr ind_number; + min_constr_number := !min_constr_number + length ; + (kn,!ind_number),1 + end + in + (* we can then build the final proof term *) + let app_constructor = applist((mkConstruct(constructor)),constructor_args) in + (* an apply the tactic *) + let res,hres = + match generate_fresh_id (id_of_string "z") (ids(* @this_branche_ids *)) 2 with + | [res;hres] -> res,hres + | _ -> assert false + in + observe (str "constructor := " ++ Printer.pr_lconstr_env (pf_env g) app_constructor); + ( + tclTHENSEQ + [ + (* unfolding of all the defined variables introduced by this branch *) + observe_tac "unfolding" pre_tac; + (* $zeta$ normalizing of the conclusion *) + h_reduce + (Rawterm.Cbv + { Rawterm.all_flags with + Rawterm.rDelta = false ; + Rawterm.rConst = [] + } + ) + onConcl; + (* introducing the the result of the graph and the equality hypothesis *) + observe_tac "introducing" (tclMAP h_intro [res;hres]); + (* replacing [res] with its value *) + observe_tac "rewriting res value" (Equality.rewriteLR (mkVar hres)); + (* Conclusion *) + observe_tac "exact" (h_exact app_constructor) + ] + ) + g + in + (* end of branche proof *) + let param_names = fst (list_chop princ_infos.nparams args_names) in + let params = List.map mkVar param_names in + let lemmas = Array.to_list (Array.map (fun c -> applist(c,params)) lemmas) in + (* The bindings of the principle + that is the params of the principle and the different lemma types + *) + let bindings = + let params_bindings,avoid = + List.fold_left2 + (fun (bindings,avoid) (x,_,_) p -> + let id = Namegen.next_ident_away (Nameops.out_name x) avoid in + (dummy_loc,Rawterm.NamedHyp id,p)::bindings,id::avoid + ) + ([],pf_ids_of_hyps g) + princ_infos.params + (List.rev params) + in + let lemmas_bindings = + List.rev (fst (List.fold_left2 + (fun (bindings,avoid) (x,_,_) p -> + let id = Namegen.next_ident_away (Nameops.out_name x) avoid in + (dummy_loc,Rawterm.NamedHyp id,(nf_zeta p))::bindings,id::avoid) + ([],avoid) + princ_infos.predicates + (lemmas))) + in + Rawterm.ExplicitBindings (params_bindings@lemmas_bindings) + in + tclTHENSEQ + [ observe_tac "intro args_names" (tclMAP h_intro args_names); + observe_tac "principle" (assert_by + (Name principle_id) + princ_type + (h_exact f_principle)); + tclTHEN_i + (observe_tac "functional_induction" ( + fun g -> + observe + (pr_constr_with_binding (Printer.pr_lconstr_env (pf_env g)) (mkVar principle_id,bindings)); + functional_induction false (applist(funs_constr.(i),List.map mkVar args_names)) + (Some (mkVar principle_id,bindings)) + pat g + )) + (fun i g -> observe_tac ("proving branche "^string_of_int i) (prove_branche i) g ) + ] + g + +(* [generalize_dependent_of x hyp g] + generalize every hypothesis which depends of [x] but [hyp] +*) +let generalize_dependent_of x hyp g = + tclMAP + (function + | (id,None,t) when not (id = hyp) && + (Termops.occur_var (pf_env g) x t) -> tclTHEN (h_generalize [mkVar id]) (thin [id]) + | _ -> tclIDTAC + ) + (pf_hyps g) + g + + + + + + (* [intros_with_rewrite] do the intros in each branch and treat each new hypothesis + (unfolding, substituting, destructing cases \ldots) + *) +let rec intros_with_rewrite g = + observe_tac "intros_with_rewrite" intros_with_rewrite_aux g +and intros_with_rewrite_aux : tactic = + fun g -> + let eq_ind = Coqlib.build_coq_eq () in + match kind_of_term (pf_concl g) with + | Prod(_,t,t') -> + begin + match kind_of_term t with + | App(eq,args) when (eq_constr eq eq_ind) -> + if Reductionops.is_conv (pf_env g) (project g) args.(1) args.(2) + then + let id = pf_get_new_id (id_of_string "y") g in + tclTHENSEQ [ h_intro id; thin [id]; intros_with_rewrite ] g + + else if isVar args.(1) + then + let id = pf_get_new_id (id_of_string "y") g in + tclTHENSEQ [ h_intro id; + generalize_dependent_of (destVar args.(1)) id; + tclTRY (Equality.rewriteLR (mkVar id)); + intros_with_rewrite + ] + g + else + begin + let id = pf_get_new_id (id_of_string "y") g in + tclTHENSEQ[ + h_intro id; + tclTRY (Equality.rewriteLR (mkVar id)); + intros_with_rewrite + ] g + end + | Ind _ when eq_constr t (Coqlib.build_coq_False ()) -> + Tauto.tauto g + | Case(_,_,v,_) -> + tclTHENSEQ[ + h_case false (v,Rawterm.NoBindings); + intros_with_rewrite + ] g + | LetIn _ -> + tclTHENSEQ[ + h_reduce + (Rawterm.Cbv + {Rawterm.all_flags + with Rawterm.rDelta = false; + }) + onConcl + ; + intros_with_rewrite + ] g + | _ -> + let id = pf_get_new_id (id_of_string "y") g in + tclTHENSEQ [ h_intro id;intros_with_rewrite] g + end + | LetIn _ -> + tclTHENSEQ[ + h_reduce + (Rawterm.Cbv + {Rawterm.all_flags + with Rawterm.rDelta = false; + }) + onConcl + ; + intros_with_rewrite + ] g + | _ -> tclIDTAC g + +let rec reflexivity_with_destruct_cases g = + let destruct_case () = + try + match kind_of_term (snd (destApp (pf_concl g))).(2) with + | Case(_,_,v,_) -> + tclTHENSEQ[ + h_case false (v,Rawterm.NoBindings); + intros; + observe_tac "reflexivity_with_destruct_cases" reflexivity_with_destruct_cases + ] + | _ -> reflexivity + with _ -> reflexivity + in + let eq_ind = Coqlib.build_coq_eq () in + let discr_inject = + Tacticals.onAllHypsAndConcl ( + fun sc g -> + match sc with + None -> tclIDTAC g + | Some id -> + match kind_of_term (pf_type_of g (mkVar id)) with + | App(eq,[|_;t1;t2|]) when eq_constr eq eq_ind -> + if Equality.discriminable (pf_env g) (project g) t1 t2 + then Equality.discrHyp id g + else if Equality.injectable (pf_env g) (project g) t1 t2 + then tclTHENSEQ [Equality.injHyp id;thin [id];intros_with_rewrite] g + else tclIDTAC g + | _ -> tclIDTAC g + ) + in + (tclFIRST + [ reflexivity; + tclTHEN (tclPROGRESS discr_inject) (destruct_case ()); + (* We reach this point ONLY if + the same value is matched (at least) two times + along binding path. + In this case, either we have a discriminable hypothesis and we are done, + either at least an injectable one and we do the injection before continuing + *) + tclTHEN (tclPROGRESS discr_inject ) reflexivity_with_destruct_cases + ]) + g + + +(* [prove_fun_complete funs graphs schemes lemmas_types_infos i] + is the tactic used to prove completness lemma. + + [funcs], [graphs] [schemes] [lemmas_types_infos] are the mutually recursive functions + (resp. definitions of the graphs of the functions, principles and correctness lemma types) to prove correct. + + [i] is the indice of the function to prove complete + + The lemma to prove if suppose to have been generated by [generate_type] (in $\zeta$ normal form that is + it looks like~: + [\forall (x_1:t_1)\ldots(x_n:t_n), forall res, + graph\ x_1\ldots x_n\ res, \rightarrow res = f x_1\ldots x_n in] + + + The sketch of the proof is the following one~: + \begin{enumerate} + \item intros until $H:graph\ x_1\ldots x_n\ res$ + \item $elim\ H$ using schemes.(i) + \item for each generated branch, intro the news hyptohesis, for each such hyptohesis [h], if [h] has + type [x=?] with [x] a variable, then subst [x], + if [h] has type [t=?] with [t] not a variable then rewrite [t] in the subterms, else + if [h] is a match then destruct it, else do just introduce it, + after all intros, the conclusion should be a reflexive equality. + \end{enumerate} + +*) + + +let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic = + fun g -> + (* We compute the types of the different mutually recursive lemmas + in $\zeta$ normal form + *) + let lemmas = + Array.map + (fun (_,(ctxt,concl)) -> nf_zeta (Termops.it_mkLambda_or_LetIn ~init:concl ctxt)) + lemmas_types_infos + in + (* We get the constant and the principle corresponding to this lemma *) + let f = funcs.(i) in + let graph_principle = nf_zeta schemes.(i) in + let princ_type = pf_type_of g graph_principle in + let princ_infos = Tactics.compute_elim_sig princ_type in + (* Then we get the number of argument of the function + and compute a fresh name for each of them + *) + let nb_fun_args = nb_prod (pf_concl g) - 2 in + let args_names = generate_fresh_id (id_of_string "x") [] nb_fun_args in + let ids = args_names@(pf_ids_of_hyps g) in + (* and fresh names for res H and the principle (cf bug bug #1174) *) + let res,hres,graph_principle_id = + match generate_fresh_id (id_of_string "z") ids 3 with + | [res;hres;graph_principle_id] -> res,hres,graph_principle_id + | _ -> assert false + in + let ids = res::hres::graph_principle_id::ids in + (* we also compute fresh names for each hyptohesis of each branche of the principle *) + let branches = List.rev princ_infos.branches in + let intro_pats = + List.map + (fun (_,_,br_type) -> + List.map + (fun id -> id) + (generate_fresh_id (id_of_string "y") ids (nb_prod br_type)) + ) + branches + in + (* We will need to change the function by its body + using [f_equation] if it is recursive (that is the graph is infinite + or unfold if the graph is finite + *) + let rewrite_tac j ids : tactic = + let graph_def = graphs.(j) in + let infos = try find_Function_infos (destConst funcs.(j)) with Not_found -> error "No graph found" in + if infos.is_general || Rtree.is_infinite graph_def.mind_recargs + then + let eq_lemma = + try Option.get (infos).equation_lemma + with Option.IsNone -> anomaly "Cannot find equation lemma" + in + tclTHENSEQ[ + tclMAP h_intro ids; + Equality.rewriteLR (mkConst eq_lemma); + (* Don't forget to $\zeta$ normlize the term since the principles have been $\zeta$-normalized *) + h_reduce + (Rawterm.Cbv + {Rawterm.all_flags + with Rawterm.rDelta = false; + }) + onConcl + ; + h_generalize (List.map mkVar ids); + thin ids + ] + else unfold_in_concl [(all_occurrences,Names.EvalConstRef (destConst f))] + in + (* The proof of each branche itself *) + let ind_number = ref 0 in + let min_constr_number = ref 0 in + let prove_branche i g = + (* we fist compute the inductive corresponding to the branch *) + let this_ind_number = + let constructor_num = i - !min_constr_number in + let length = Array.length (graphs.(!ind_number).Declarations.mind_consnames) in + if constructor_num <= length + then !ind_number + else + begin + incr ind_number; + min_constr_number := !min_constr_number + length; + !ind_number + end + in + let this_branche_ids = List.nth intro_pats (pred i) in + tclTHENSEQ[ + (* we expand the definition of the function *) + observe_tac "rewrite_tac" (rewrite_tac this_ind_number this_branche_ids); + (* introduce hypothesis with some rewrite *) + observe_tac "intros_with_rewrite" intros_with_rewrite; + (* The proof is (almost) complete *) + observe_tac "reflexivity" (reflexivity_with_destruct_cases) + ] + g + in + let params_names = fst (list_chop princ_infos.nparams args_names) in + let params = List.map mkVar params_names in + tclTHENSEQ + [ tclMAP h_intro (args_names@[res;hres]); + observe_tac "h_generalize" + (h_generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)]); + h_intro graph_principle_id; + observe_tac "" (tclTHEN_i + (observe_tac "elim" ((elim false (mkVar hres,Rawterm.NoBindings) (Some (mkVar graph_principle_id,Rawterm.NoBindings))))) + (fun i g -> observe_tac "prove_branche" (prove_branche i) g )) + ] + g + + + + +let do_save () = Lemmas.save_named false + + +(* [derive_correctness make_scheme functional_induction funs graphs] create correctness and completeness + lemmas for each function in [funs] w.r.t. [graphs] + + [make_scheme] is Functional_principle_types.make_scheme (dependency pb) and + [functional_induction] is Indfun.functional_induction (same pb) +*) + +let derive_correctness make_scheme functional_induction (funs: constant list) (graphs:inductive list) = + let funs = Array.of_list funs and graphs = Array.of_list graphs in + let funs_constr = Array.map mkConst funs in + try + let graphs_constr = Array.map mkInd graphs in + let lemmas_types_infos = + Util.array_map2_i + (fun i f_constr graph -> + let const_of_f = destConst f_constr in + let (type_of_lemma_ctxt,type_of_lemma_concl) as type_info = + generate_type false const_of_f graph i + in + let type_of_lemma = Termops.it_mkProd_or_LetIn ~init:type_of_lemma_concl type_of_lemma_ctxt in + let type_of_lemma = nf_zeta type_of_lemma in + observe (str "type_of_lemma := " ++ Printer.pr_lconstr type_of_lemma); + type_of_lemma,type_info + ) + funs_constr + graphs_constr + in + let schemes = + (* The functional induction schemes are computed and not saved if there is more that one function + if the block contains only one function we can safely reuse [f_rect] + *) + try + if Array.length funs_constr <> 1 then raise Not_found; + [| find_induction_principle funs_constr.(0) |] + with Not_found -> + Array.of_list + (List.map + (fun entry -> + (entry.Entries.const_entry_body, Option.get entry.Entries.const_entry_type ) + ) + (make_scheme (array_map_to_list (fun const -> const,Rawterm.RType None) funs)) + ) + in + let proving_tac = + prove_fun_correct functional_induction funs_constr graphs_constr schemes lemmas_types_infos + in + Array.iteri + (fun i f_as_constant -> + let f_id = id_of_label (con_label f_as_constant) in + Lemmas.start_proof + (*i The next call to mk_correct_id is valid since we are constructing the lemma + Ensures by: obvious + i*) + (mk_correct_id f_id) + (Decl_kinds.Global,(Decl_kinds.Proof Decl_kinds.Theorem)) + (fst lemmas_types_infos.(i)) + (fun _ _ -> ()); + Pfedit.by (observe_tac ("prove correctness ("^(string_of_id f_id)^")") (proving_tac i)); + do_save (); + let finfo = find_Function_infos f_as_constant in + update_Function + {finfo with + correctness_lemma = Some (destConst (Tacinterp.constr_of_id (Global.env ())(mk_correct_id f_id))) + } + + ) + funs; + let lemmas_types_infos = + Util.array_map2_i + (fun i f_constr graph -> + let const_of_f = destConst f_constr in + let (type_of_lemma_ctxt,type_of_lemma_concl) as type_info = + generate_type true const_of_f graph i + in + let type_of_lemma = Termops.it_mkProd_or_LetIn ~init:type_of_lemma_concl type_of_lemma_ctxt in + let type_of_lemma = nf_zeta type_of_lemma in + observe (str "type_of_lemma := " ++ Printer.pr_lconstr type_of_lemma); + type_of_lemma,type_info + ) + funs_constr + graphs_constr + in + let kn,_ as graph_ind = destInd graphs_constr.(0) in + let mib,mip = Global.lookup_inductive graph_ind in + let schemes = + Array.of_list + (Indrec.build_mutual_induction_scheme (Global.env ()) Evd.empty + (Array.to_list + (Array.mapi + (fun i _ -> (kn,i),true,InType) + mib.Declarations.mind_packets + ) + ) + ) + in + let proving_tac = + prove_fun_complete funs_constr mib.Declarations.mind_packets schemes lemmas_types_infos + in + Array.iteri + (fun i f_as_constant -> + let f_id = id_of_label (con_label f_as_constant) in + Lemmas.start_proof + (*i The next call to mk_complete_id is valid since we are constructing the lemma + Ensures by: obvious + i*) + (mk_complete_id f_id) + (Decl_kinds.Global,(Decl_kinds.Proof Decl_kinds.Theorem)) + (fst lemmas_types_infos.(i)) + (fun _ _ -> ()); + Pfedit.by (observe_tac ("prove completeness ("^(string_of_id f_id)^")") (proving_tac i)); + do_save (); + let finfo = find_Function_infos f_as_constant in + update_Function + {finfo with + completeness_lemma = Some (destConst (Tacinterp.constr_of_id (Global.env ())(mk_complete_id f_id))) + } + ) + funs; + with e -> + (* In case of problem, we reset all the lemmas *) + (*i The next call to mk_correct_id is valid since we are erasing the lemmas + Ensures by: obvious + i*) + let first_lemma_id = + let f_id = id_of_label (con_label funs.(0)) in + + mk_correct_id f_id + in + ignore(try Vernacentries.vernac_reset_name (Util.dummy_loc,first_lemma_id) with _ -> ()); + raise e + + + + + +(***********************************************) + +(* [revert_graph kn post_tac hid] transforme an hypothesis [hid] having type Ind(kn,num) t1 ... tn res + when [kn] denotes a graph block into + f_num t1... tn = res (by applying [f_complete] to the first type) before apply post_tac on the result + + if the type of hypothesis has not this form or if we cannot find the completeness lemma then we do nothing +*) +let revert_graph kn post_tac hid g = + let typ = pf_type_of g (mkVar hid) in + match kind_of_term typ with + | App(i,args) when isInd i -> + let ((kn',num) as ind') = destInd i in + if kn = kn' + then (* We have generated a graph hypothesis so that we must change it if we can *) + let info = + try find_Function_of_graph ind' + with Not_found -> (* The graphs are mutually recursive but we cannot find one of them !*) + anomaly "Cannot retrieve infos about a mutual block" + in + (* if we can find a completeness lemma for this function + then we can come back to the functional form. If not, we do nothing + *) + match info.completeness_lemma with + | None -> tclIDTAC g + | Some f_complete -> + let f_args,res = array_chop (Array.length args - 1) args in + tclTHENSEQ + [ + h_generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])]; + thin [hid]; + h_intro hid; + post_tac hid + ] + g + + else tclIDTAC g + | _ -> tclIDTAC g + + +(* + [functional_inversion hid fconst f_correct ] is the functional version of [inversion] + + [hid] is the hypothesis to invert, [fconst] is the function to invert and [f_correct] + is the correctness lemma for [fconst]. + + The sketch is the follwing~: + \begin{enumerate} + \item Transforms the hypothesis [hid] such that its type is now $res\ =\ f\ t_1 \ldots t_n$ + (fails if it is not possible) + \item replace [hid] with $R\_f t_1 \ldots t_n res$ using [f_correct] + \item apply [inversion] on [hid] + \item finally in each branch, replace each hypothesis [R\_f ..] by [f ...] using [f_complete] (whenever + such a lemma exists) + \end{enumerate} +*) + +let functional_inversion kn hid fconst f_correct : tactic = + fun g -> + let old_ids = List.fold_right Idset.add (pf_ids_of_hyps g) Idset.empty in + let type_of_h = pf_type_of g (mkVar hid) in + match kind_of_term type_of_h with + | App(eq,args) when eq_constr eq (Coqlib.build_coq_eq ()) -> + let pre_tac,f_args,res = + match kind_of_term args.(1),kind_of_term args.(2) with + | App(f,f_args),_ when eq_constr f fconst -> + ((fun hid -> h_symmetry (onHyp hid)),f_args,args.(2)) + |_,App(f,f_args) when eq_constr f fconst -> + ((fun hid -> tclIDTAC),f_args,args.(1)) + | _ -> (fun hid -> tclFAIL 1 (mt ())),[||],args.(2) + in + tclTHENSEQ[ + pre_tac hid; + h_generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])]; + thin [hid]; + h_intro hid; + Inv.inv FullInversion None (Rawterm.NamedHyp hid); + (fun g -> + let new_ids = List.filter (fun id -> not (Idset.mem id old_ids)) (pf_ids_of_hyps g) in + tclMAP (revert_graph kn pre_tac) (hid::new_ids) g + ); + ] g + | _ -> tclFAIL 1 (mt ()) g + + + +let invfun qhyp f = + let f = + match f with + | ConstRef f -> f + | _ -> raise (Util.UserError("",str "Not a function")) + in + try + let finfos = find_Function_infos f in + let f_correct = mkConst(Option.get finfos.correctness_lemma) + and kn = fst finfos.graph_ind + in + Tactics.try_intros_until (fun hid -> functional_inversion kn hid (mkConst f) f_correct) qhyp + with + | Not_found -> error "No graph found" + | Option.IsNone -> error "Cannot use equivalence with graph!" + + +let invfun qhyp f g = + match f with + | Some f -> invfun qhyp f g + | None -> + Tactics.try_intros_until + (fun hid g -> + let hyp_typ = pf_type_of g (mkVar hid) in + match kind_of_term hyp_typ with + | App(eq,args) when eq_constr eq (Coqlib.build_coq_eq ()) -> + begin + let f1,_ = decompose_app args.(1) in + try + if not (isConst f1) then failwith ""; + let finfos = find_Function_infos (destConst f1) in + let f_correct = mkConst(Option.get finfos.correctness_lemma) + and kn = fst finfos.graph_ind + in + functional_inversion kn hid f1 f_correct g + with | Failure "" | Option.IsNone | Not_found -> + try + let f2,_ = decompose_app args.(2) in + if not (isConst f2) then failwith ""; + let finfos = find_Function_infos (destConst f2) in + let f_correct = mkConst(Option.get finfos.correctness_lemma) + and kn = fst finfos.graph_ind + in + functional_inversion kn hid f2 f_correct g + with + | Failure "" -> + errorlabstrm "" (str "Hypothesis" ++ Ppconstr.pr_id hid ++ str " must contain at leat one Function") + | Option.IsNone -> + if do_observe () + then + error "Cannot use equivalence with graph for any side of the equality" + else errorlabstrm "" (str "Cannot find inversion information for hypothesis " ++ Ppconstr.pr_id hid) + | Not_found -> + if do_observe () + then + error "No graph found for any side of equality" + else errorlabstrm "" (str "Cannot find inversion information for hypothesis " ++ Ppconstr.pr_id hid) + end + | _ -> errorlabstrm "" (Ppconstr.pr_id hid ++ str " must be an equality ") + ) + qhyp + g diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml new file mode 100644 index 00000000..f596e2d7 --- /dev/null +++ b/plugins/funind/merge.ml @@ -0,0 +1,1032 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* false) t1 t2 + then true + else false + +let rec compare_constr' t1 t2 = + if compare_constr_nosub t1 t2 + then true + else (compare_constr (compare_constr') t1 t2) + +let rec substitterm prof t by_t in_u = + if (compare_constr' (lift prof t) in_u) + then (lift prof by_t) + else map_constr_with_binders succ + (fun i -> substitterm i t by_t) prof in_u + +let lift_ldecl n ldecl = List.map (fun (x,y) -> x,lift n y) ldecl + +let understand = Pretyping.Default.understand Evd.empty (Global.env()) + +(** Operations on names and identifiers *) +let id_of_name = function + Anonymous -> id_of_string "H" + | Name id -> id;; +let name_of_string str = Name (id_of_string str) +let string_of_name nme = string_of_id (id_of_name nme) + +(** [isVarf f x] returns [true] if term [x] is of the form [(Var f)]. *) +let isVarf f x = + match x with + | RVar (_,x) -> Pervasives.compare x f = 0 + | _ -> false + +(** [ident_global_exist id] returns true if identifier [id] is linked + in global environment. *) +let ident_global_exist id = + try + let ans = CRef (Libnames.Ident (dummy_loc,id)) in + let _ = ignore (Constrintern.intern_constr Evd.empty (Global.env()) ans) in + true + with _ -> false + +(** [next_ident_fresh id] returns a fresh identifier (ie not linked in + global env) with base [id]. *) +let next_ident_fresh (id:identifier) = + let res = ref id in + while ident_global_exist !res do res := Nameops.lift_subscript !res done; + !res + + +(** {2 Debugging} *) +(* comment this line to see debug msgs *) +let msg x = () ;; let pr_lconstr c = str "" +(* uncomment this to see debugging *) +let prconstr c = msg (str" " ++ Printer.pr_lconstr c) +let prconstrnl c = msg (str" " ++ Printer.pr_lconstr c ++ str"\n") +let prlistconstr lc = List.iter prconstr lc +let prstr s = msg(str s) +let prNamedConstr s c = + begin + msg(str ""); + msg(str(s^" {§ ") ++ Printer.pr_lconstr c ++ str " §} "); + msg(str ""); + end +let prNamedRConstr s c = + begin + msg(str ""); + msg(str(s^" {§ ") ++ Printer.pr_rawconstr c ++ str " §} "); + msg(str ""); + end +let prNamedLConstr_aux lc = List.iter (prNamedConstr "\n") lc +let prNamedLConstr s lc = + begin + prstr "[§§§ "; + prstr s; + prNamedLConstr_aux lc; + prstr " §§§]\n"; + end +let prNamedLDecl s lc = + begin + prstr s; prstr "\n"; + List.iter (fun (nm,_,tp) -> prNamedConstr (string_of_name nm) tp) lc; + prstr "\n"; + end +let prNamedRLDecl s lc = + begin + prstr s; prstr "\n"; prstr "{§§ "; + List.iter + (fun x -> + match x with + | (nm,None,Some tp) -> prNamedRConstr (string_of_name nm) tp + | (nm,Some bdy,None) -> prNamedRConstr ("(letin) "^string_of_name nm) bdy + | _ -> assert false + ) lc; + prstr " §§}\n"; + prstr "\n"; + end + +let showind (id:identifier) = + let cstrid = Tacinterp.constr_of_id (Global.env()) id in + let ind1,cstrlist = Inductiveops.find_inductive (Global.env()) Evd.empty cstrid in + let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) ind1 in + List.iter (fun (nm, optcstr, tp) -> + print_string (string_of_name nm^":"); + prconstr tp; print_string "\n") + ib1.mind_arity_ctxt; + (match ib1.mind_arity with + | Monomorphic x -> + Printf.printf "arity :"; prconstr x.mind_user_arity + | Polymorphic x -> + Printf.printf "arity : universe?"); + Array.iteri + (fun i x -> Printf.printf"type constr %d :" i ; prconstr x) + ib1.mind_user_lc + +(** {2 Misc} *) + +exception Found of int + +(* Array scanning *) + +let array_prfx (arr: 'a array) (pred: int -> 'a -> bool): int = + try + for i=0 to Array.length arr - 1 do if pred i (arr.(i)) then raise (Found i) done; + Array.length arr (* all elt are positive *) + with Found i -> i + +let array_fold_lefti (f: int -> 'a -> 'b -> 'a) (acc:'a) (arr:'b array): 'a = + let i = ref 0 in + Array.fold_left + (fun acc x -> + let res = f !i acc x in i := !i + 1; res) + acc arr + +(* Like list_chop but except that [i] is the size of the suffix of [l]. *) +let list_chop_end i l = + let size_prefix = List.length l -i in + if size_prefix < 0 then failwith "list_chop_end" + else list_chop size_prefix l + +let list_fold_lefti (f: int -> 'a -> 'b -> 'a) (acc:'a) (arr:'b list): 'a = + let i = ref 0 in + List.fold_left + (fun acc x -> + let res = f !i acc x in i := !i + 1; res) + acc arr + +let list_filteri (f: int -> 'a -> bool) (l:'a list):'a list = + let i = ref 0 in + List.filter (fun x -> let res = f !i x in i := !i + 1; res) l + + +(** Iteration module *) +module For = +struct + let rec map i j (f: int -> 'a) = if i>j then [] else f i :: (map (i+1) j f) + let rec foldup i j (f: 'a -> int -> 'a) acc = + if i>j then acc else let newacc = f acc i in foldup (i+1) j f newacc + let rec folddown i j (f: 'a -> int -> 'a) acc = + if i>j then acc else let newacc = f acc j in folddown i (j-1) f newacc + let fold i j = if i Printf.sprintf "Linked %d" i + | Unlinked -> Printf.sprintf "Unlinked" + | Funres -> Printf.sprintf "Funres" + +let linkmonad f lnkvar = + match lnkvar with + | Linked i -> Linked (f i) + | Unlinked -> Unlinked + | Funres -> Funres + +let linklift lnkvar i = linkmonad (fun x -> x+i) lnkvar + +(* This map is used to deal with debruijn linked indices. *) +module Link = Map.Make (struct type t = int let compare = Pervasives.compare end) + +let pr_links l = + Printf.printf "links:\n"; + Link.iter (fun k e -> Printf.printf "%d : %s\n" k (prlinked e)) l; + Printf.printf "_____________\n" + +type 'a merged_arg = + | Prm_stable of 'a + | Prm_linked of 'a + | Prm_arg of 'a + | Arg_stable of 'a + | Arg_linked of 'a + | Arg_funres + +(** Information about graph merging of two inductives. + All rel_decl list are IN REVERSE ORDER (ie well suited for compose) *) + +type merge_infos = + { + ident:identifier; (** new inductive name *) + mib1: mutual_inductive_body; + oib1: one_inductive_body; + mib2: mutual_inductive_body; + oib2: one_inductive_body; + + (** Array of links of the first inductive (should be all stable) *) + lnk1: int merged_arg array; + + (** Array of links of the second inductive (point to the first ind param/args) *) + lnk2: int merged_arg array; + + (** rec params which remain rec param (ie not linked) *) + recprms1: rel_declaration list; + recprms2: rel_declaration list; + nrecprms1: int; + nrecprms2: int; + + (** rec parms which became non parm (either linked to something + or because after a rec parm that became non parm) *) + otherprms1: rel_declaration list; + otherprms2: rel_declaration list; + notherprms1:int; + notherprms2:int; + + (** args which remain args in merge *) + args1:rel_declaration list; + args2:rel_declaration list; + nargs1:int; + nargs2:int; + + (** functional result args *) + funresprms1: rel_declaration list; + funresprms2: rel_declaration list; + nfunresprms1:int; + nfunresprms2:int; + } + + +let pr_merginfo x = + let i,s= + match x with + | Prm_linked i -> Some i,"Prm_linked" + | Arg_linked i -> Some i,"Arg_linked" + | Prm_stable i -> Some i,"Prm_stable" + | Prm_arg i -> Some i,"Prm_arg" + | Arg_stable i -> Some i,"Arg_stable" + | Arg_funres -> None , "Arg_funres" in + match i with + | Some i -> Printf.sprintf "%s(%d)" s i + | None -> Printf.sprintf "%s" s + +let isPrm_stable x = match x with Prm_stable _ -> true | _ -> false + +(* ?? prm_linked?? *) +let isArg_stable x = match x with Arg_stable _ | Prm_arg _ -> true | _ -> false + +let is_stable x = + match x with Arg_stable _ | Prm_stable _ | Prm_arg _ -> true | _ -> false + +let isArg_funres x = match x with Arg_funres -> true | _ -> false + +let filter_shift_stable (lnk:int merged_arg array) (l:'a list): 'a list = + let prms = list_filteri (fun i _ -> isPrm_stable lnk.(i)) l in + let args = list_filteri (fun i _ -> isArg_stable lnk.(i)) l in + let fres = list_filteri (fun i _ -> isArg_funres lnk.(i)) l in + prms@args@fres + +(** Reverse the link map, keeping only linked vars, elements are list + of int as several vars may be linked to the same var. *) +let revlinked lnk = + For.fold 0 (Array.length lnk - 1) + (fun acc k -> + match lnk.(k) with + | Unlinked | Funres -> acc + | Linked i -> + let old = try Link.find i acc with Not_found -> [] in + Link.add i (k::old) acc) + Link.empty + +let array_switch arr i j = + let aux = arr.(j) in arr.(j) <- arr.(i); arr.(i) <- aux + +let filter_shift_stable_right (lnk:int merged_arg array) (l:'a list): 'a list = + let larr = Array.of_list l in + let _ = + Array.iteri + (fun j x -> + match x with + | Prm_linked i -> array_switch larr i j + | Arg_linked i -> array_switch larr i j + | Prm_stable i -> () + | Prm_arg i -> () + | Arg_stable i -> () + | Arg_funres -> () + ) lnk in + filter_shift_stable lnk (Array.to_list larr) + + + + +(** {1 Utilities for merging} *) + +let ind1name = id_of_string "__ind1" +let ind2name = id_of_string "__ind2" + +(** Performs verifications on two graphs before merging: they must not + be co-inductive, and for the moment they must not be mutual + either. *) +let verify_inds mib1 mib2 = + if not mib1.mind_finite then error "First argument is coinductive"; + if not mib2.mind_finite then error "Second argument is coinductive"; + if mib1.mind_ntypes <> 1 then error "First argument is mutual"; + if mib2.mind_ntypes <> 1 then error "Second argument is mutual"; + () + +(* +(** [build_raw_params prms_decl avoid] returns a list of variables + attributed to the list of decl [prms_decl], avoiding names in + [avoid]. *) +let build_raw_params prms_decl avoid = + let dummy_constr = compose_prod (List.map (fun (x,_,z) -> x,z) prms_decl) (mkRel 1) in + let _ = prNamedConstr "DUMMY" dummy_constr in + let dummy_rawconstr = Detyping.detype false avoid [] dummy_constr in + let _ = prNamedRConstr "RAWDUMMY" dummy_rawconstr in + let res,_ = raw_decompose_prod dummy_rawconstr in + let comblist = List.combine prms_decl res in + comblist, res , (avoid @ (Idset.elements (ids_of_rawterm dummy_rawconstr))) +*) + +let ids_of_rawlist avoid rawl = + List.fold_left Idset.union avoid (List.map ids_of_rawterm rawl) + + + +(** {1 Merging function graphs} *) + +(** [shift_linked_params mib1 mib2 lnk] Computes which parameters (rec + uniform and ordinary ones) of mutual inductives [mib1] and [mib2] + remain uniform when linked by [lnk]. All parameters are + considered, ie we take parameters of the first inductive body of + [mib1] and [mib2]. + + Explanation: The two inductives have parameters, some of the first + are recursively uniform, some of the last are functional result of + the functional graph. + + (I x1 x2 ... xk ... xk' ... xn) + (J y1 y2 ... xl ... yl' ... ym) + + Problem is, if some rec unif params are linked to non rec unif + ones, they become non rec (and the following too). And functinal + argument have to be shifted at the end *) +let shift_linked_params mib1 mib2 (lnk1:linked_var array) (lnk2:linked_var array) id = + let _ = prstr "\nYOUHOU shift\n" in + let linked_targets = revlinked lnk2 in + let is_param_of_mib1 x = x < mib1.mind_nparams_rec in + let is_param_of_mib2 x = x < mib2.mind_nparams_rec in + let is_targetted_by_non_recparam_lnk1 i = + try + let targets = Link.find i linked_targets in + List.exists (fun x -> not (is_param_of_mib2 x)) targets + with Not_found -> false in + let mlnk1 = + Array.mapi + (fun i lkv -> + let isprm = is_param_of_mib1 i in + let prmlost = is_targetted_by_non_recparam_lnk1 i in + match isprm , prmlost, lnk1.(i) with + | true , true , _ -> Prm_arg i (* recparam becoming ordinary *) + | true , false , _-> Prm_stable i (* recparam remains recparam*) + | false , false , Funres -> Arg_funres + | _ , _ , Funres -> assert false (* fun res cannot be a rec param or lost *) + | false , _ , _ -> Arg_stable i) (* Args of lnk1 are not linked *) + lnk1 in + let mlnk2 = + Array.mapi + (fun i lkv -> + (* Is this correct if some param of ind2 is lost? *) + let isprm = is_param_of_mib2 i in + match isprm , lnk2.(i) with + | true , Linked j when not (is_param_of_mib1 j) -> + Prm_arg j (* recparam becoming ordinary *) + | true , Linked j -> Prm_linked j (*recparam linked to recparam*) + | true , Unlinked -> Prm_stable i (* recparam remains recparam*) + | false , Linked j -> Arg_linked j (* Args of lnk2 lost *) + | false , Unlinked -> Arg_stable i (* Args of lnk2 remains *) + | false , Funres -> Arg_funres + | true , Funres -> assert false (* fun res cannot be a rec param *) + ) + lnk2 in + let oib1 = mib1.mind_packets.(0) in + let oib2 = mib2.mind_packets.(0) in + (* count params remaining params *) + let n_params1 = array_prfx mlnk1 (fun i x -> not (isPrm_stable x)) in + let n_params2 = array_prfx mlnk2 (fun i x -> not (isPrm_stable x)) in + let bldprms arity_ctxt mlnk = + list_fold_lefti + (fun i (acc1,acc2,acc3,acc4) x -> + prstr (pr_merginfo mlnk.(i));prstr "\n"; + match mlnk.(i) with + | Prm_stable _ -> x::acc1 , acc2 , acc3, acc4 + | Prm_arg _ -> acc1 , x::acc2 , acc3, acc4 + | Arg_stable _ -> acc1 , acc2 , x::acc3, acc4 + | Arg_funres -> acc1 , acc2 , acc3, x::acc4 + | _ -> acc1 , acc2 , acc3, acc4) + ([],[],[],[]) arity_ctxt in +(* let arity_ctxt2 = + build_raw_params oib2.mind_arity_ctxt + (Idset.elements (ids_of_rawterm oib1.mind_arity_ctxt)) in*) + let recprms1,otherprms1,args1,funresprms1 = bldprms (List.rev oib1.mind_arity_ctxt) mlnk1 in + let _ = prstr "\n\n\n" in + let recprms2,otherprms2,args2,funresprms2 = bldprms (List.rev oib2.mind_arity_ctxt) mlnk2 in + let _ = prstr "\notherprms1:\n" in + let _ = + List.iter (fun (x,_,y) -> prstr (string_of_name x^" : ");prconstr y;prstr "\n") + otherprms1 in + let _ = prstr "\notherprms2:\n" in + let _ = + List.iter (fun (x,_,y) -> prstr (string_of_name x^" : ");prconstr y;prstr "\n") + otherprms2 in + { + ident=id; + mib1=mib1; + oib1 = oib1; + mib2=mib2; + oib2 = oib2; + lnk1 = mlnk1; + lnk2 = mlnk2; + nrecprms1 = n_params1; + recprms1 = recprms1; + otherprms1 = otherprms1; + args1 = args1; + funresprms1 = funresprms1; + notherprms1 = Array.length mlnk1 - n_params1; + nfunresprms1 = List.length funresprms1; + nargs1 = List.length args1; + nrecprms2 = n_params2; + recprms2 = recprms2; + otherprms2 = otherprms2; + args2 = args2; + funresprms2 = funresprms2; + notherprms2 = Array.length mlnk2 - n_params2; + nargs2 = List.length args2; + nfunresprms2 = List.length funresprms2; + } + + + + +(** {1 Merging functions} *) + +exception NoMerge + +let rec merge_app c1 c2 id1 id2 shift filter_shift_stable = + let lnk = Array.append shift.lnk1 shift.lnk2 in + match c1 , c2 with + | RApp(_,f1, arr1), RApp(_,f2,arr2) when isVarf id1 f1 && isVarf id2 f2 -> + let _ = prstr "\nICI1!\n";Pp.flush_all() in + let args = filter_shift_stable lnk (arr1 @ arr2) in + RApp (dummy_loc,RVar (dummy_loc,shift.ident) , args) + | RApp(_,f1, arr1), RApp(_,f2,arr2) -> raise NoMerge + | RLetIn(_,nme,bdy,trm) , _ -> + let _ = prstr "\nICI2!\n";Pp.flush_all() in + let newtrm = merge_app trm c2 id1 id2 shift filter_shift_stable in + RLetIn(dummy_loc,nme,bdy,newtrm) + | _, RLetIn(_,nme,bdy,trm) -> + let _ = prstr "\nICI3!\n";Pp.flush_all() in + let newtrm = merge_app c1 trm id1 id2 shift filter_shift_stable in + RLetIn(dummy_loc,nme,bdy,newtrm) + | _ -> let _ = prstr "\nICI4!\n";Pp.flush_all() in + raise NoMerge + +let rec merge_app_unsafe c1 c2 shift filter_shift_stable = + let lnk = Array.append shift.lnk1 shift.lnk2 in + match c1 , c2 with + | RApp(_,f1, arr1), RApp(_,f2,arr2) -> + let args = filter_shift_stable lnk (arr1 @ arr2) in + RApp (dummy_loc,RVar(dummy_loc,shift.ident) , args) + (* FIXME: what if the function appears in the body of the let? *) + | RLetIn(_,nme,bdy,trm) , _ -> + let _ = prstr "\nICI2 '!\n";Pp.flush_all() in + let newtrm = merge_app_unsafe trm c2 shift filter_shift_stable in + RLetIn(dummy_loc,nme,bdy,newtrm) + | _, RLetIn(_,nme,bdy,trm) -> + let _ = prstr "\nICI3 '!\n";Pp.flush_all() in + let newtrm = merge_app_unsafe c1 trm shift filter_shift_stable in + RLetIn(dummy_loc,nme,bdy,newtrm) + | _ -> let _ = prstr "\nICI4 '!\n";Pp.flush_all() in raise NoMerge + + + +(* Heuristic when merging two lists of hypothesis: merge every rec + calls of branch 1 with all rec calls of branch 2. *) +(* TODO: reecrire cette heuristique (jusqu'a merge_types) *) +let rec merge_rec_hyps shift accrec + (ltyp:(Names.name * rawconstr option * rawconstr option) list) + filter_shift_stable : (Names.name * rawconstr option * rawconstr option) list = + let mergeonehyp t reldecl = + match reldecl with + | (nme,x,Some (RApp(_,i,args) as ind)) + -> nme,x, Some (merge_app_unsafe ind t shift filter_shift_stable) + | (nme,Some _,None) -> error "letins with recursive calls not treated yet" + | (nme,None,Some _) -> assert false + | (nme,None,None) | (nme,Some _,Some _) -> assert false in + match ltyp with + | [] -> [] + | (nme,None,Some (RApp(_,f, largs) as t)) :: lt when isVarf ind2name f -> + let rechyps = List.map (mergeonehyp t) accrec in + rechyps @ merge_rec_hyps shift accrec lt filter_shift_stable + | e::lt -> e :: merge_rec_hyps shift accrec lt filter_shift_stable + + +let rec build_suppl_reccall (accrec:(name * rawconstr) list) concl2 shift = + List.map (fun (nm,tp) -> (nm,merge_app_unsafe tp concl2 shift)) accrec + + +let find_app (nme:identifier) ltyp = + try + ignore + (List.map + (fun x -> + match x with + | _,None,Some (RApp(_,f,_)) when isVarf nme f -> raise (Found 0) + | _ -> ()) + ltyp); + false + with Found _ -> true + +let prnt_prod_or_letin nm letbdy typ = + match letbdy , typ with + | Some lbdy , None -> prNamedRConstr ("(letin) " ^ string_of_name nm) lbdy + | None , Some tp -> prNamedRConstr (string_of_name nm) tp + | _ , _ -> assert false + + +let rec merge_types shift accrec1 + (ltyp1:(name * rawconstr option * rawconstr option) list) + (concl1:rawconstr) (ltyp2:(name * rawconstr option * rawconstr option) list) concl2 + : (name * rawconstr option * rawconstr option) list * rawconstr = + let _ = prstr "MERGE_TYPES\n" in + let _ = prstr "ltyp 1 : " in + let _ = List.iter (fun (nm,lbdy,tp) -> prnt_prod_or_letin nm lbdy tp) ltyp1 in + let _ = prstr "\nltyp 2 : " in + let _ = List.iter (fun (nm,lbdy,tp) -> prnt_prod_or_letin nm lbdy tp) ltyp2 in + let _ = prstr "\n" in + let res = + match ltyp1 with + | [] -> + let isrec1 = (accrec1<>[]) in + let isrec2 = find_app ind2name ltyp2 in + let rechyps = + if isrec1 && isrec2 + then (* merge_rec_hyps shift accrec1 ltyp2 filter_shift_stable *) + merge_rec_hyps shift [name_of_string "concl1",None,Some concl1] ltyp2 + filter_shift_stable_right + @ merge_rec_hyps shift accrec1 [name_of_string "concl2",None, Some concl2] + filter_shift_stable + else if isrec1 + (* if rec calls in accrec1 and not in ltyp2, add one to ltyp2 *) + then + merge_rec_hyps shift accrec1 + (ltyp2@[name_of_string "concl2",None,Some concl2]) filter_shift_stable + else if isrec2 + then merge_rec_hyps shift [name_of_string "concl1",None,Some concl1] ltyp2 + filter_shift_stable_right + else ltyp2 in + let _ = prstr"\nrechyps : " in + let _ = List.iter(fun (nm,lbdy,tp)-> prnt_prod_or_letin nm lbdy tp) rechyps in + let _ = prstr "MERGE CONCL : " in + let _ = prNamedRConstr "concl1" concl1 in + let _ = prstr " with " in + let _ = prNamedRConstr "concl2" concl2 in + let _ = prstr "\n" in + let concl = + merge_app concl1 concl2 ind1name ind2name shift filter_shift_stable in + let _ = prstr "FIN " in + let _ = prNamedRConstr "concl" concl in + let _ = prstr "\n" in + + rechyps , concl + | (nme,None, Some t1)as e ::lt1 -> + (match t1 with + | RApp(_,f,carr) when isVarf ind1name f -> + merge_types shift (e::accrec1) lt1 concl1 ltyp2 concl2 + | _ -> + let recres, recconcl2 = + merge_types shift accrec1 lt1 concl1 ltyp2 concl2 in + ((nme,None,Some t1) :: recres) , recconcl2) + | (nme,Some bd, None) ::lt1 -> + (* FIXME: what if ind1name appears in bd? *) + let recres, recconcl2 = + merge_types shift accrec1 lt1 concl1 ltyp2 concl2 in + ((nme,Some bd,None) :: recres) , recconcl2 + | (_,None,None)::_ | (_,Some _,Some _)::_ -> assert false + in + res + + +(** [build_link_map_aux allargs1 allargs2 shift] returns the mapping of + linked args [allargs2] to target args of [allargs1] as specified + in [shift]. [allargs1] and [allargs2] are in reverse order. Also + returns the list of unlinked vars of [allargs2]. *) +let build_link_map_aux (allargs1:identifier array) (allargs2:identifier array) + (lnk:int merged_arg array) = + array_fold_lefti + (fun i acc e -> + if i = Array.length lnk - 1 then acc (* functional arg, not in allargs *) + else + match e with + | Prm_linked j | Arg_linked j -> Idmap.add allargs2.(i) allargs1.(j) acc + | _ -> acc) + Idmap.empty lnk + +let build_link_map allargs1 allargs2 lnk = + let allargs1 = + Array.of_list (List.rev (List.map (fun (x,_,_) -> id_of_name x) allargs1)) in + let allargs2 = + Array.of_list (List.rev (List.map (fun (x,_,_) -> id_of_name x) allargs2)) in + build_link_map_aux allargs1 allargs2 lnk + + +(** [merge_one_constructor lnk shift typcstr1 typcstr2] merges the two + constructor rawtypes [typcstr1] and [typcstr2]. [typcstr1] and + [typcstr2] contain all parameters (including rec. unif. ones) of + their inductive. + + if [typcstr1] and [typcstr2] are of the form: + + forall recparams1, forall ordparams1, H1a -> H2a... (I1 x1 y1 ... z1) + forall recparams2, forall ordparams2, H2b -> H2b... (I2 x2 y2 ... z2) + + we build: + + forall recparams1 (recparams2 without linked params), + forall ordparams1 (ordparams2 without linked params), + H1a' -> H2a' -> ... -> H2a' -> H2b'(shifted) -> ... + -> (newI x1 ... z1 x2 y2 ...z2 without linked params) + + where Hix' have been adapted, ie: + - linked vars have been changed, + - rec calls to I1 and I2 have been replaced by rec calls to + newI. More precisely calls to I1 and I2 have been merge by an + experimental heuristic (in particular if n o rec calls for I1 + or I2 is found, we use the conclusion as a rec call). See + [merge_types] above. + + Precond: vars sets of [typcstr1] and [typcstr2] must be disjoint. + + TODO: return nothing if equalities (after linking) are contradictory. *) +let merge_one_constructor (shift:merge_infos) (typcstr1:rawconstr) + (typcstr2:rawconstr) : rawconstr = + (* FIXME: les noms des parametres corerspondent en principe au + parametres du niveau mib, mais il faudrait s'en assurer *) + (* shift.nfunresprmsx last args are functional result *) + let nargs1 = + shift.mib1.mind_nparams + shift.oib1.mind_nrealargs - shift.nfunresprms1 in + let nargs2 = + shift.mib2.mind_nparams + shift.oib2.mind_nrealargs - shift.nfunresprms2 in + let allargs1,rest1 = raw_decompose_prod_or_letin_n nargs1 typcstr1 in + let allargs2,rest2 = raw_decompose_prod_or_letin_n nargs2 typcstr2 in + (* Build map of linked args of [typcstr2], and apply it to [typcstr2]. *) + let linked_map = build_link_map allargs1 allargs2 shift.lnk2 in + let rest2 = change_vars linked_map rest2 in + let hyps1,concl1 = raw_decompose_prod_or_letin rest1 in + let hyps2,concl2' = raw_decompose_prod_or_letin rest2 in + let ltyp,concl2 = + merge_types shift [] (List.rev hyps1) concl1 (List.rev hyps2) concl2' in + let _ = prNamedRLDecl "ltyp result:" ltyp in + let typ = raw_compose_prod_or_letin concl2 (List.rev ltyp) in + let revargs1 = + list_filteri (fun i _ -> isArg_stable shift.lnk1.(i)) (List.rev allargs1) in + let _ = prNamedRLDecl "ltyp allargs1" allargs1 in + let _ = prNamedRLDecl "ltyp revargs1" revargs1 in + let revargs2 = + list_filteri (fun i _ -> isArg_stable shift.lnk2.(i)) (List.rev allargs2) in + let _ = prNamedRLDecl "ltyp allargs2" allargs2 in + let _ = prNamedRLDecl "ltyp revargs2" revargs2 in + let typwithprms = + raw_compose_prod_or_letin typ (List.rev revargs2 @ List.rev revargs1) in + typwithprms + + +(** constructor numbering *) +let fresh_cstror_suffix , cstror_suffix_init = + let cstror_num = ref 0 in + (fun () -> + let res = string_of_int !cstror_num in + cstror_num := !cstror_num + 1; + res) , + (fun () -> cstror_num := 0) + +(** [merge_constructor_id id1 id2 shift] returns the identifier of the + new constructor from the id of the two merged constructor and + the merging info. *) +let merge_constructor_id id1 id2 shift:identifier = + let id = string_of_id shift.ident ^ "_" ^ fresh_cstror_suffix () in + next_ident_fresh (id_of_string id) + + + +(** [merge_constructors lnk shift avoid] merges the two list of + constructor [(name*type)]. These are translated to rawterms + first, each of them having distinct var names. *) +let rec merge_constructors (shift:merge_infos) (avoid:Idset.t) + (typcstr1:(identifier * rawconstr) list) + (typcstr2:(identifier * rawconstr) list) : (identifier * rawconstr) list = + List.flatten + (List.map + (fun (id1,rawtyp1) -> + List.map + (fun (id2,rawtyp2) -> + let typ = merge_one_constructor shift rawtyp1 rawtyp2 in + let newcstror_id = merge_constructor_id id1 id2 shift in + let _ = prstr "\n**************\n" in + newcstror_id , typ) + typcstr2) + typcstr1) + +(** [merge_inductive_body lnk shift avoid oib1 oib2] merges two + inductive bodies [oib1] and [oib2], linking with [lnk], params + info in [shift], avoiding identifiers in [avoid]. *) +let rec merge_inductive_body (shift:merge_infos) avoid (oib1:one_inductive_body) + (oib2:one_inductive_body) = + (* building rawconstr type of constructors *) + let mkrawcor nme avoid typ = + (* first replace rel 1 by a varname *) + let substindtyp = substitterm 0 (mkRel 1) (mkVar nme) typ in + Detyping.detype false (Idset.elements avoid) [] substindtyp in + let lcstr1: rawconstr list = + Array.to_list (Array.map (mkrawcor ind1name avoid) oib1.mind_user_lc) in + (* add to avoid all indentifiers of lcstr1 *) + let avoid2 = Idset.union avoid (ids_of_rawlist avoid lcstr1) in + let lcstr2 = + Array.to_list (Array.map (mkrawcor ind2name avoid2) oib2.mind_user_lc) in + let avoid3 = Idset.union avoid (ids_of_rawlist avoid lcstr2) in + + let params1 = + try fst (raw_decompose_prod_n shift.nrecprms1 (List.hd lcstr1)) + with _ -> [] in + let params2 = + try fst (raw_decompose_prod_n shift.nrecprms2 (List.hd lcstr2)) + with _ -> [] in + + let lcstr1 = List.combine (Array.to_list oib1.mind_consnames) lcstr1 in + let lcstr2 = List.combine (Array.to_list oib2.mind_consnames) lcstr2 in + + cstror_suffix_init(); + params1,params2,merge_constructors shift avoid3 lcstr1 lcstr2 + + +(** [merge_mutual_inductive_body lnk mib1 mib2 shift] merge mutual + inductive bodies [mib1] and [mib2] linking vars with + [lnk]. [shift] information on parameters of the new inductive. + For the moment, inductives are supposed to be non mutual. +*) +let rec merge_mutual_inductive_body + (mib1:mutual_inductive_body) (mib2:mutual_inductive_body) (shift:merge_infos) = + (* Mutual not treated, we take first ind body of each. *) + merge_inductive_body shift Idset.empty mib1.mind_packets.(0) mib2.mind_packets.(0) + + +let rawterm_to_constr_expr x = (* build a constr_expr from a rawconstr *) + Flags.with_option Flags.raw_print (Constrextern.extern_rawtype Idset.empty) x + +let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) = + let params = prms2 @ prms1 in + let resparams = + List.fold_left + (fun acc (nme,tp) -> + let _ = prstr "param :" in + let _ = prNamedRConstr (string_of_name nme) tp in + let _ = prstr " ; " in + let typ = rawterm_to_constr_expr tp in + LocalRawAssum ([(dummy_loc,nme)], Topconstr.default_binder_kind, typ) :: acc) + [] params in + let concl = Constrextern.extern_constr false (Global.env()) concl in + let arity,_ = + List.fold_left + (fun (acc,env) (nm,_,c) -> + let typ = Constrextern.extern_constr false env c in + let newenv = Environ.push_rel (nm,None,c) env in + CProdN (dummy_loc, [[(dummy_loc,nm)],Topconstr.default_binder_kind,typ] , acc) , newenv) + (concl,Global.env()) + (shift.funresprms2 @ shift.funresprms1 + @ shift.args2 @ shift.args1 @ shift.otherprms2 @ shift.otherprms1) in + resparams,arity + + + +(** [rawterm_list_to_inductive_expr ident rawlist] returns the + induct_expr corresponding to the the list of constructor types + [rawlist], named ident. + FIXME: params et cstr_expr (arity) *) +let rawterm_list_to_inductive_expr prms1 prms2 mib1 mib2 shift + (rawlist:(identifier * rawconstr) list) = + let lident = dummy_loc, shift.ident in + let bindlist , cstr_expr = (* params , arities *) + merge_rec_params_and_arity prms1 prms2 shift mkSet in + let lcstor_expr : (bool * (lident * constr_expr)) list = + List.map (* zeta_normalize t ? *) + (fun (id,t) -> false, ((dummy_loc,id),rawterm_to_constr_expr t)) + rawlist in + lident , bindlist , Some cstr_expr , lcstor_expr + + + +let mkProd_reldecl (rdecl:rel_declaration) (t2:rawconstr) = + match rdecl with + | (nme,None,t) -> + let traw = Detyping.detype false [] [] t in + RProd (dummy_loc,nme,Explicit,traw,t2) + | (_,Some _,_) -> assert false + + + + +let mkProd_reldecl (rdecl:rel_declaration) (t2:rawconstr) = + match rdecl with + | (nme,None,t) -> + let traw = Detyping.detype false [] [] t in + RProd (dummy_loc,nme,Explicit,traw,t2) + | (_,Some _,_) -> assert false + + +(** [merge_inductive ind1 ind2 lnk] merges two graphs, linking + variables specified in [lnk]. Graphs are not supposed to be mutual + inductives for the moment. *) +let merge_inductive (ind1: inductive) (ind2: inductive) + (lnk1: linked_var array) (lnk2: linked_var array) id = + let env = Global.env() in + let mib1,_ = Inductive.lookup_mind_specif env ind1 in + let mib2,_ = Inductive.lookup_mind_specif env ind2 in + let _ = verify_inds mib1 mib2 in (* raises an exception if something wrong *) + (* compute params that become ordinary args (because linked to ord. args) *) + let shift_prm = shift_linked_params mib1 mib2 lnk1 lnk2 id in + let prms1,prms2, rawlist = merge_mutual_inductive_body mib1 mib2 shift_prm in + let _ = prstr "\nrawlist : " in + let _ = + List.iter (fun (nm,tp) -> prNamedRConstr (string_of_id nm) tp;prstr "\n") rawlist in + let _ = prstr "\nend rawlist\n" in +(* FIX: retransformer en constr ici + let shift_prm = + { shift_prm with + recprms1=prms1; + recprms1=prms1; + } in *) + let indexpr = rawterm_list_to_inductive_expr prms1 prms2 mib1 mib2 shift_prm rawlist in + (* Declare inductive *) + let indl,_,_ = Command.extract_mutual_inductive_declaration_components [(indexpr,[])] in + let mie,impls = Command.interp_mutual_inductive indl [] true (* means: not coinductive *) in + (* Declare the mutual inductive block with its associated schemes *) + ignore (Command.declare_mutual_inductive_with_eliminations Declare.UserVerbose mie impls) + + +(* Find infos on identifier id. *) +let find_Function_infos_safe (id:identifier): Indfun_common.function_info = + let kn_of_id x = + let f_ref = Libnames.Ident (dummy_loc,x) in + locate_with_msg (str "Don't know what to do with " ++ Libnames.pr_reference f_ref) + locate_constant f_ref in + try find_Function_infos (kn_of_id id) + with Not_found -> + errorlabstrm "indfun" (Nameops.pr_id id ++ str " has no functional scheme") + +(** [merge id1 id2 args1 args2 id] builds and declares a new inductive + type called [id], representing the merged graphs of both graphs + [ind1] and [ind2]. identifiers occuring in both arrays [args1] and + [args2] are considered linked (i.e. are the same variable) in the + new graph. + + Warning: For the moment, repetitions of an id in [args1] or + [args2] are not supported. *) +let merge (id1:identifier) (id2:identifier) (args1:identifier array) + (args2:identifier array) id : unit = + let finfo1 = find_Function_infos_safe id1 in + let finfo2 = find_Function_infos_safe id2 in + (* FIXME? args1 are supposed unlinked. mergescheme (G x x) ?? *) + (* We add one arg (functional arg of the graph) *) + let lnk1 = Array.make (Array.length args1 + 1) Unlinked in + let lnk2' = (* args2 may be linked to args1 members. FIXME: same + as above: vars may be linked inside args2?? *) + Array.mapi + (fun i c -> + match array_find_i (fun i x -> x=c) args1 with + | Some j -> Linked j + | None -> Unlinked) + args2 in + (* We add one arg (functional arg of the graph) *) + let lnk2 = Array.append lnk2' (Array.make 1 Unlinked) in + (* setting functional results *) + let _ = lnk1.(Array.length lnk1 - 1) <- Funres in + let _ = lnk2.(Array.length lnk2 - 1) <- Funres in + merge_inductive finfo1.graph_ind finfo2.graph_ind lnk1 lnk2 id + + +let remove_last_arg c = + let (x,y) = decompose_prod c in + let xnolast = List.rev (List.tl (List.rev x)) in + compose_prod xnolast y + +let rec remove_n_fst_list n l = if n=0 then l else remove_n_fst_list (n-1) (List.tl l) +let remove_n_last_list n l = List.rev (remove_n_fst_list n (List.rev l)) + +let remove_last_n_arg n c = + let (x,y) = decompose_prod c in + let xnolast = remove_n_last_list n x in + compose_prod xnolast y + +(* [funify_branches relinfo nfuns branch] returns the branch [branch] + of the relinfo [relinfo] modified to fit in a functional principle. + Things to do: + - remove indargs from rel applications + - replace *variables only* corresponding to function (recursive) + results by the actual function application. *) +let funify_branches relinfo nfuns branch = + let mut_induct, induct = + match relinfo.indref with + | None -> assert false + | Some (IndRef ((mutual_ind,i) as ind)) -> mutual_ind,ind + | _ -> assert false in + let is_dom c = + match kind_of_term c with + | Ind((u,_)) | Construct((u,_),_) -> u = mut_induct + | _ -> false in + let _dom_i c = + assert (is_dom c); + match kind_of_term c with + | Ind((u,i)) | Construct((u,_),i) -> i + | _ -> assert false in + let _is_pred c shift = + match kind_of_term c with + | Rel i -> let reali = i-shift in (reali>=0 && reali false in + (* FIXME: *) + (Anonymous,Some mkProp,mkProp) + + +let relprinctype_to_funprinctype relprinctype nfuns = + let relinfo = compute_elim_sig relprinctype in + assert (not relinfo.farg_in_concl); + assert (relinfo.indarg_in_concl); + (* first remove indarg and indarg_in_concl *) + let relinfo_noindarg = { relinfo with + indarg_in_concl = false; indarg = None; + concl = remove_last_arg (pop relinfo.concl); } in + (* the nfuns last induction arguments are functional ones: remove them *) + let relinfo_argsok = { relinfo_noindarg with + nargs = relinfo_noindarg.nargs - nfuns; + (* args is in reverse order, so remove fst *) + args = remove_n_fst_list nfuns relinfo_noindarg.args; + concl = popn nfuns relinfo_noindarg.concl + } in + let new_branches = + List.map (funify_branches relinfo_argsok nfuns) relinfo_argsok.branches in + let relinfo_branches = { relinfo_argsok with branches = new_branches } in + relinfo_branches + +(* @article{ bundy93rippling, + author = "Alan Bundy and Andrew Stevens and Frank van Harmelen and Andrew Ireland and Alan Smaill", + title = "Rippling: A Heuristic for Guiding Inductive Proofs", + journal = "Artificial Intelligence", + volume = "62", + number = "2", + pages = "185-253", + year = "1993", + url = "citeseer.ist.psu.edu/bundy93rippling.html" } + + *) +(* +*** Local Variables: *** +*** compile-command: "make -C ../.. plugins/funind/merge.cmo" *** +*** indent-tabs-mode: nil *** +*** End: *** +*) diff --git a/plugins/funind/rawterm_to_relation.ml b/plugins/funind/rawterm_to_relation.ml new file mode 100644 index 00000000..3c3a36f0 --- /dev/null +++ b/plugins/funind/rawterm_to_relation.ml @@ -0,0 +1,1419 @@ +open Printer +open Pp +open Names +open Term +open Rawterm +open Libnames +open Indfun_common +open Util +open Rawtermops + +let observe strm = + if do_observe () + then Pp.msgnl strm + else () +let observennl strm = + if do_observe () + then Pp.msg strm + else () + + +type binder_type = + | Lambda of name + | Prod of name + | LetIn of name + +type raw_context = (binder_type*rawconstr) list + +(* + compose_raw_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 + binders corresponding to the bt_i's +*) +let compose_raw_context = + let compose_binder (bt,t) acc = + match bt with + | Lambda n -> mkRLambda(n,t,acc) + | Prod n -> mkRProd(n,t,acc) + | LetIn n -> mkRLetIn(n,t,acc) + in + List.fold_right compose_binder + + +(* + The main part deals with building a list of raw constructor expressions + from the rhs of a fixpoint equation. +*) + +type 'a build_entry_pre_return = + { + context : raw_context; (* the binding context of the result *) + value : 'a; (* The value *) + } + +type 'a build_entry_return = + { + result : 'a build_entry_pre_return list; + to_avoid : identifier list + } + +(* + [combine_results combine_fun res1 res2] combine two results [res1] and [res2] + w.r.t. [combine_fun]. + + Informally, both [res1] and [res2] are lists of "constructors" [res1_1;...] + and [res2_1,....] and we need to produce + [combine_fun res1_1 res2_1;combine_fun res1_1 res2_2;........] +*) + +let combine_results + (combine_fun : 'a build_entry_pre_return -> 'b build_entry_pre_return -> + 'c build_entry_pre_return + ) + (res1: 'a build_entry_return) + (res2 : 'b build_entry_return) + : 'c build_entry_return + = + 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 + ) + res1.result + in (* and then we flatten the map *) + { + result = List.concat pre_result; + to_avoid = list_union res1.to_avoid res2.to_avoid + } + + +(* + The combination function for an argument with a list of argument +*) + +let combine_args arg args = + { + context = arg.context@args.context; + (* Note that the binding context of [arg] MUST be placed before the one of + [args] in order to preserve possible type dependencies + *) + value = arg.value::args.value; + } + + +let ids_of_binder = function + | LetIn Anonymous | Prod Anonymous | Lambda Anonymous -> [] + | LetIn (Name id) | Prod (Name id) | Lambda (Name id) -> [id] + +let rec change_vars_in_binder mapping = function + [] -> [] + | (bt,t)::l -> + let new_mapping = List.fold_right Idmap.remove (ids_of_binder bt) mapping in + (bt,change_vars mapping t):: + (if idmap_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 List.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 = List.append (ids_of_binder bt) + +let apply_args ctxt body args = + let need_convert_id avoid id = + List.exists (is_free_in id) args || List.mem id avoid + in + let need_convert avoid bt = + List.exists (need_convert_id avoid) (ids_of_binder bt) + in + let next_name_away (na:name) (mapping: identifier Idmap.t) (avoid: identifier list) = + match na with + | Name id when List.mem id avoid -> + let new_id = Namegen.next_ident_away id avoid in + Name new_id,Idmap.add id new_id mapping,new_id::avoid + | _ -> na,mapping,avoid + in + let next_bt_away bt (avoid:identifier list) = + match bt with + | LetIn na -> + let new_na,mapping,new_avoid = next_name_away na Idmap.empty avoid in + LetIn new_na,mapping,new_avoid + | Prod na -> + let new_na,mapping,new_avoid = next_name_away na Idmap.empty avoid in + Prod new_na,mapping,new_avoid + | Lambda na -> + let new_na,mapping,new_avoid = next_name_away na Idmap.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) + | [],_ -> (* no more fun *) + let f,args' = raw_decompose_app body in + (ctxt,mkRApp(f,args'@args)) + | (Lambda Anonymous,t)::ctxt',arg::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::avoid in + let new_id = Namegen.next_ident_away id new_avoid in + let new_avoid' = new_id :: new_avoid in + let mapping = Idmap.add id new_id Idmap.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::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 + in + do_apply [] ctxt body args + + +let combine_app f args = + let new_ctxt,new_value = apply_args f.context f.value args.value in + { + (* Note that the binding context of [args] MUST be placed before the one of + the applied value in order to preserve possible type dependencies + *) + context = args.context@new_ctxt; + value = new_value; + } + +let combine_lam n t b = + { + context = []; + value = mkRLambda(n, compose_raw_context t.context t.value, + compose_raw_context b.context b.value ) + } + + + +let combine_prod n t b = + { context = t.context@((Prod n,t.value)::b.context); value = b.value} + +let combine_letin n t b = + { context = t.context@((LetIn n,t.value)::b.context); value = b.value} + + +let mk_result ctxt value avoid = + { + result = + [{context = ctxt; + value = value}] + ; + to_avoid = avoid + } +(************************************************* + Some functions to deal with overlapping patterns +**************************************************) + +let coq_True_ref = + lazy (Coqlib.gen_reference "" ["Init";"Logic"] "True") + +let coq_False_ref = + lazy (Coqlib.gen_reference "" ["Init";"Logic"] "False") + +(* + [make_discr_match_el \[e1,...en\]] builds match e1,...,en with + (the list of expresions on which we will do the matching) + *) +let make_discr_match_el = + List.map (fun e -> (e,(Anonymous,None))) + +(* + [make_discr_match_brl i \[pat_1,...,pat_n\]] constructs a discrimination pattern matching on the ith expression. + that is. + match ?????? with \\ + | pat_1 => False \\ + | pat_{i-1} => False \\ + | pat_i => True \\ + | pat_{i+1} => False \\ + \vdots + | pat_n => False + end +*) +let make_discr_match_brl i = + list_map_i + (fun j (_,idl,patl,_) -> + if j=i + then (dummy_loc,idl,patl, mkRRef (Lazy.force coq_True_ref)) + else (dummy_loc,idl,patl, mkRRef (Lazy.force coq_False_ref)) + ) + 0 +(* + [make_discr_match brl el i] generates an hypothesis such that it reduce to true iff + brl_{i} is the first branch matched by [el] + + Used when we want to simulate the coq pattern matching algorithm +*) +let make_discr_match brl = + fun el i -> + mkRCases(None, + make_discr_match_el el, + make_discr_match_brl i brl) + +let pr_name = function + | Name id -> Ppconstr.pr_id id + | Anonymous -> str "_" + +(**********************************************************************) +(* functions used to build case expression from lettuple and if ones *) +(**********************************************************************) + +(* [build_constructors_of_type] construct the array of pattern of its inductive argument*) +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 constructref = ConstructRef(construct) in + let _implicit_positions_of_cst = + Impargs.implicits_of_global constructref + in + let cst_narg = + Inductiveops.mis_constructor_nargs_env + (Global.env ()) + construct + in + let argl = + if argl = [] + then + Array.to_list + (Array.init (cst_narg - npar) (fun _ -> mkRHole ()) + ) + else argl + in + let pat_as_term = + mkRApp(mkRRef (ConstructRef(ind',i+1)),argl) + in + cases_pattern_of_rawconstr Anonymous pat_as_term + ) + ind.Declarations.mind_consnames + +(* [find_type_of] very naive attempts to discover the type of an if or a letin *) +let rec find_type_of nb b = + let f,_ = raw_decompose_app b in + match f with + | RRef(_,ref) -> + begin + let ind_type = + match ref with + | VarRef _ | ConstRef _ -> + let constr_of_ref = constr_of_global ref in + let type_of_ref = Typing.type_of (Global.env ()) Evd.empty constr_of_ref in + let (_,ret_type) = Reduction.dest_prod (Global.env ()) type_of_ref in + let ret_type,_ = decompose_app ret_type in + if not (isInd ret_type) then + begin +(* Pp.msgnl (str "not an inductive" ++ pr_lconstr ret_type); *) + raise (Invalid_argument "not an inductive") + end; + destInd ret_type + | IndRef ind -> ind + | ConstructRef c -> fst c + in + let _,ind_type_info = Inductive.lookup_mind_specif (Global.env()) ind_type in + if not (Array.length ind_type_info.Declarations.mind_consnames = nb ) + then raise (Invalid_argument "find_type_of : not a valid inductive"); + ind_type + end + | RCast(_,b,_) -> find_type_of nb b + | RApp _ -> assert false (* we have decomposed any application via raw_decompose_app *) + | _ -> raise (Invalid_argument "not a ref") + + + + +(******************) +(* Main functions *) +(******************) + + + +let raw_push_named (na,raw_value,raw_typ) env = + match na with + | Anonymous -> env + | Name id -> + let value = Option.map (Pretyping.Default.understand Evd.empty env) raw_value in + let typ = Pretyping.Default.understand_type Evd.empty env raw_typ in + Environ.push_named (id,value,typ) env + + +let add_pat_variables pat typ env : Environ.env = + let rec add_pat_variables env pat typ : Environ.env = + observe (str "new rel env := " ++ Printer.pr_rel_context_of env); + + match pat with + | PatVar(_,na) -> Environ.push_rel (na,None,typ) env + | PatCstr(_,c,patl,na) -> + let Inductiveops.IndType(indf,indargs) = + try Inductiveops.find_rectype env Evd.empty typ + with Not_found -> assert false + in + let constructors = Inductiveops.get_constructors env indf in + let constructor : Inductiveops.constructor_summary = List.find (fun cs -> cs.Inductiveops.cs_cstr = c) (Array.to_list constructors) in + let cs_args_types :types list = List.map (fun (_,_,t) -> t) 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 ( + Sign.fold_rel_context + (fun (na,v,t) (env,ctxt) -> + match na with + | Anonymous -> assert false + | Name id -> + let new_t = substl ctxt t in + let new_v = Option.map (substl ctxt) v in + observe (str "for variable " ++ Ppconstr.pr_id id ++ fnl () ++ + str "old type := " ++ Printer.pr_lconstr t ++ fnl () ++ + str "new type := " ++ Printer.pr_lconstr new_t ++ fnl () ++ + Option.fold_right (fun v _ -> str "old value := " ++ Printer.pr_lconstr v ++ fnl ()) v (mt ()) ++ + Option.fold_right (fun v _ -> str "new value := " ++ Printer.pr_lconstr v ++ fnl ()) new_v (mt ()) + ); + (Environ.push_named (id,new_v,new_t) env,mkVar id::ctxt) + ) + (Environ.rel_context new_env) + ~init:(env,[]) + ) + in + observe (str "new var env := " ++ Printer.pr_named_context_of res); + res + + + + +let rec pattern_to_term_and_type env typ = function + | PatVar(loc,Anonymous) -> assert false + | PatVar(loc,Name id) -> + mkRVar id + | PatCstr(loc,constr,patternl,_) -> + let cst_narg = + Inductiveops.mis_constructor_nargs_env + (Global.env ()) + constr + in + let Inductiveops.IndType(indf,indargs) = + try Inductiveops.find_rectype env Evd.empty typ + with Not_found -> assert false + in + let constructors = Inductiveops.get_constructors env indf in + let constructor = List.find (fun cs -> cs.Inductiveops.cs_cstr = constr) (Array.to_list constructors) in + let cs_args_types :types list = List.map (fun (_,_,t) -> t) constructor.Inductiveops.cs_args in + 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 false [] (Termops.names_of_rel_context env) csta.(i)) + ) + in + let patl_as_term = + List.map2 (pattern_to_term_and_type env) (List.rev cs_args_types) patternl + in + mkRApp(mkRRef(ConstructRef constr), + implicit_args@patl_as_term + ) + +(* [build_entry_lc funnames avoid rt] construct the list (in fact a build_entry_return) + of constructors corresponding to [rt] when replacing calls to [funnames] by calls to the + corresponding graphs. + + + The idea to transform a term [t] into a list of constructors [lc] is the following: + \begin{itemize} + \item if the term is a binder (bind x, body) then first compute [lc'] the list corresponding + to [body] and add (bind x. _) to each elements of [lc] + \item if the term has the form (g t1 ... ... tn) where g does not appears in (fnames) + then compute [lc1] ... [lcn] the lists of constructors corresponding to [t1] ... [tn], + then combine those lists and [g] as follows~: for each element [c1,...,cn] of [lc1\times...\times lcn], + [g c1 ... cn] is an element of [lc] + \item if the term has the form (f t1 .... tn) where [f] appears in [fnames] then + compute [lc1] ... [lcn] the lists of constructors corresponding to [t1] ... [tn], + then compute those lists and [f] as follows~: for each element [c1,...,cn] of [lc1\times...\times lcn] + create a new variable [res] and [forall res, R_f c1 ... cn res] is in [lc] + \item if the term is a cast just treat its body part + \item + if the term is a match, an if or a lettuple then compute the lists corresponding to each branch of the case + and concatenate them (informally, each branch of a match produces a new constructor) + \end{itemize} + + WARNING: The terms constructed here are only USING the rawconstr syntax but are highly bad formed. + We must wait to have complete all the current calculi to set the recursive calls. + At this point, each term [f t1 ... tn] (where f appears in [funnames]) is replaced by + a pseudo term [forall res, res t1 ... tn, res]. A reconstruction phase is done later. + We in fact not create a constructor list since then end of each constructor has not the expected form + but only the value of the function +*) + + +let rec build_entry_lc env funnames avoid rt : rawconstr build_entry_return = + observe (str " Entering : " ++ Printer.pr_rawconstr rt); + match rt with + | RRef _ | RVar _ | REvar _ | RPatVar _ | RSort _ | RHole _ -> + (* do nothing (except changing type of course) *) + mk_result [] rt avoid + | RApp(_,_,_) -> + let f,args = raw_decompose_app rt in + let args_res : (rawconstr 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 funnames ctxt_argsl.to_avoid arg in + combine_results combine_args arg_res ctxt_argsl + ) + args + (mk_result [] [] avoid) + in + begin + match f with + | RLambda _ -> + let rec aux t l = + match l with + | [] -> t + | u::l -> + match t with + | RLambda(loc,na,_,nat,b) -> + RLetIn(dummy_loc,na,u,aux b l) + | _ -> + RApp(dummy_loc,t,l) + in + build_entry_lc env funnames avoid (aux f args) + | RVar(_,id) when Idset.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 = Pretyping.Default.understand Evd.empty env rt in + let rt_typ = Typing.type_of env Evd.empty rt_as_constr in + let res_raw_type = Detyping.detype false [] (Termops.names_of_rel_context 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 = mkRVar res in + let new_result = + List.map + (fun arg_res -> + let new_hyps = + [Prod (Name res),res_raw_type; + Prod Anonymous,mkRApp(res_rt,(mkRVar 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 } + | RVar _ | REvar _ | RPatVar _ | RHole _ | RSort _ | RRef _ -> + (* 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 = mkRApp(f,args_res.value)}) + args_res.result + } + | RApp _ -> assert false (* we have collected all the app in [raw_decompose_app] *) + | RLetIn(_,n,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 avoid in + let new_avoid = id:: avoid in + let new_b = + replace_var_by_term + id + (RVar(dummy_loc,id)) + b + in + (Name new_id,new_b,new_avoid) + | _ -> n,b,avoid + in + build_entry_lc + env + funnames + avoid + (mkRLetIn(new_n,t,mkRApp(new_b,args))) + | RCases _ | RIf _ | RLetTuple _ -> + (* 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 funnames args_res.to_avoid f in + combine_results combine_app f_res args_res + | RDynamic _ ->error "Not handled RDynamic" + | RCast(_,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 + *) + build_entry_lc env funnames avoid (mkRApp(b,args)) + | RRec _ -> error "Not handled RRec" + | RProd _ -> error "Cannot apply a type" + end (* end of the application treatement *) + + | RLambda(_,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 + *) + let t_res = build_entry_lc env 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 b_res = build_entry_lc new_env funnames avoid b in + combine_results (combine_lam new_n) t_res b_res + | RProd(_,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 + *) + let t_res = build_entry_lc env funnames avoid t in + let new_env = raw_push_named (n,None,t) env in + let b_res = build_entry_lc new_env funnames avoid b in + combine_results (combine_prod n) t_res b_res + | RLetIn(_,n,v,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 + *) + let v_res = build_entry_lc env funnames avoid v in + let v_as_constr = Pretyping.Default.understand Evd.empty env v in + let v_type = Typing.type_of env Evd.empty v_as_constr in + let new_env = + match n with + Anonymous -> env + | Name id -> Environ.push_named (id,Some v_as_constr,v_type) env + in + let b_res = build_entry_lc new_env funnames avoid b in + combine_results (combine_letin n) v_res b_res + | RCases(_,_,_,el,brl) -> + (* we create the discrimination function + and treat the case itself + *) + let make_discr = make_discr_match brl in + build_entry_lc_from_case env funnames make_discr el brl avoid + | RIf(_,b,(na,e_option),lhs,rhs) -> + let b_as_constr = Pretyping.Default.understand Evd.empty env b in + let b_typ = Typing.type_of env Evd.empty b_as_constr in + let (ind,_) = + try Inductiveops.find_inductive env Evd.empty b_typ + with Not_found -> + errorlabstrm "" (str "Cannot find the inductive associated to " ++ + Printer.pr_rawconstr b ++ str " in " ++ + Printer.pr_rawconstr rt ++ str ". try again with a cast") + in + let case_pats = build_constructors_of_type ind [] in + assert (Array.length case_pats = 2); + let brl = + list_map_i + (fun i x -> (dummy_loc,[],[case_pats.(i)],x)) + 0 + [lhs;rhs] + in + let match_expr = + mkRCases(None,[(b,(Anonymous,None))],brl) + in + (* Pp.msgnl (str "new case := " ++ Printer.pr_rawconstr match_expr); *) + build_entry_lc env funnames avoid match_expr + | RLetTuple(_,nal,_,b,e) -> + begin + let nal_as_rawconstr = + List.map + (function + Name id -> mkRVar id + | Anonymous -> mkRHole () + ) + nal + in + let b_as_constr = Pretyping.Default.understand Evd.empty env b in + let b_typ = Typing.type_of env Evd.empty b_as_constr in + let (ind,_) = + try Inductiveops.find_inductive env Evd.empty b_typ + with Not_found -> + errorlabstrm "" (str "Cannot find the inductive associated to " ++ + Printer.pr_rawconstr b ++ str " in " ++ + Printer.pr_rawconstr rt ++ str ". try again with a cast") + in + let case_pats = build_constructors_of_type ind nal_as_rawconstr in + assert (Array.length case_pats = 1); + let br = + (dummy_loc,[],[case_pats.(0)],e) + in + let match_expr = mkRCases(None,[b,(Anonymous,None)],[br]) in + build_entry_lc env funnames avoid match_expr + + end + | RRec _ -> error "Not handled RRec" + | RCast(_,b,_) -> + build_entry_lc env funnames avoid b + | RDynamic _ -> error "Not handled RDynamic" +and build_entry_lc_from_case env funname make_discr + (el:tomatch_tuples) + (brl:Rawterm.cases_clauses) avoid : + rawconstr build_entry_return = + match el with + | [] -> assert false (* this case correspond to match 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 elemeent 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 funname 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 = Pretyping.Default.understand Evd.empty env case_arg in + Typing.type_of env Evd.empty 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 + env 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 acc r.to_avoid) [] results + } + +and build_entry_lc_from_case_term env types funname make_discr patterns_to_prevent brl avoid + matched_expr = + match brl with + | [] -> (* computed_branches *) {result = [];to_avoid = avoid} + | br::brl' -> + (* alpha convertion to prevent name clashes *) + let _,idl,patl,return = alpha_br avoid br in + let new_avoid = idl@avoid in (* for now we can no more use idl as an indentifier *) + (* 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 patl types env in + let not_those_patterns : (identifier list -> rawconstr -> rawconstr) 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 pat typ new_env in + List.fold_right + (fun id acc -> + let typ_of_id = + Typing.type_of env_with_pat_ids Evd.empty (mkVar id) + in + let raw_typ_of_id = + Detyping.detype false [] + (Termops.names_of_rel_context env_with_pat_ids) typ_of_id + in + mkRProd (Name id,raw_typ_of_id,acc)) + pat_ids + (raw_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 + 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 = Detyping.detype false [] (Termops.names_of_rel_context new_env) typ_as_constr in + let pat_as_term = pattern_to_term pat in + List.fold_right + (fun id acc -> + if Idset.mem id this_pat_ids + then (Prod (Name id), + let typ_of_id = Typing.type_of new_env Evd.empty (mkVar id) in + let raw_typ_of_id = + Detyping.detype false [] (Termops.names_of_rel_context new_env) typ_of_id + in + raw_typ_of_id + )::acc + else acc + ) + idl + [(Prod Anonymous,raw_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 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 } + + +let is_res id = + try + String.sub (string_of_id id) 0 3 = "res" + with Invalid_argument _ -> false + + +exception Continue +(* + The second phase which reconstruct the real type of the constructor. + rebuild the raw constructors expression. + eliminates some meaningless equalities, applies some rewrites...... +*) +let rec rebuild_cons env nb_args relname args crossed_types depth rt = + observe (str "rebuilding : " ++ pr_rawconstr rt); + + match rt with + | RProd(_,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 t with + | RApp(_,(RVar(_,res_id) as res_rt),args') when is_res res_id -> + begin + match args' with + | (RVar(_,this_relname))::args' -> + (*i The next call to mk_rel_id is + valid since we are constructing the graph + Ensures by: obvious + i*) + + let new_t = + mkRApp(mkRVar(mk_rel_id this_relname),args'@[res_rt]) + in + let t' = Pretyping.Default.understand Evd.empty env new_t in + let new_env = Environ.push_rel (n,None,t') env in + let new_b,id_to_exclude = + rebuild_cons new_env + nb_args relname + args new_crossed_types + (depth + 1) b + in + mkRProd(n,new_t,new_b), + Idset.filter not_free_in_t id_to_exclude + | _ -> (* the first args is the name of the function! *) + assert false + end + | RApp(loc1,RRef(loc2,eq_as_ref),[ty;RVar(loc3,id);rt]) + when eq_as_ref = Lazy.force Coqlib.coq_eq_ref && n = Anonymous + -> + begin + try + observe (str "computing new type for eq : " ++ pr_rawconstr rt); + let t' = + try Pretyping.Default.understand Evd.empty env t with _ -> 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 new_env = Environ.push_rel (n,None,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 + mkRProd(n,t,new_b),id_to_exclude + with Continue -> + let jmeq = Libnames.IndRef (destInd (jmeq ())) in + let ty' = Pretyping.Default.understand Evd.empty env ty in + let ind,args' = Inductive.find_inductive env ty' in + let mib,_ = Global.lookup_inductive ind in + let nparam = mib.Declarations.mind_nparams in + let params,arg' = + ((Util.list_chop nparam args')) + in + let rt_typ = + RApp(Util.dummy_loc, + RRef (Util.dummy_loc,Libnames.IndRef ind), + (List.map + (fun p -> Detyping.detype false [] + (Termops.names_of_rel_context env) + p) params)@(Array.to_list + (Array.make + (List.length args' - nparam) + (mkRHole ())))) + in + let eq' = + RApp(loc1,RRef(loc2,jmeq),[ty;RVar(loc3,id);rt_typ;rt]) + in + observe (str "computing new type for jmeq : " ++ pr_rawconstr eq'); + let eq'_as_constr = Pretyping.Default.understand Evd.empty env eq' in + observe (str " computing new type for jmeq : done") ; + let new_args = + match kind_of_term eq'_as_constr with + | App(_,[|_;_;ty;_|]) -> + let ty = Array.to_list (snd (destApp 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,_,_) = + Environ.lookup_rel (destRel var_as_constr) env + in + match na with + | Anonymous -> acc + | Name id' -> + (id',Detyping.detype false [] + (Termops.names_of_rel_context env) + arg)::acc + else if isVar var_as_constr + then (destVar var_as_constr,Detyping.detype false [] + (Termops.names_of_rel_context 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' = Pretyping.Default.understand Evd.empty env eq' in + Environ.push_rel (n,None,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 + mkRProd(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 + mkRProd(n,t,new_b),id_to_exclude + else new_b, Idset.add id id_to_exclude + *) + | _ -> + observe (str "computing new type for prod : " ++ pr_rawconstr rt); + let t' = Pretyping.Default.understand Evd.empty env t in + let new_env = Environ.push_rel (n,None,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 Idset.mem id id_to_exclude && depth >= nb_args -> + new_b,Idset.remove id + (Idset.filter not_free_in_t id_to_exclude) + | _ -> mkRProd(n,t,new_b),Idset.filter not_free_in_t id_to_exclude + end + | RLambda(_,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 + observe (str "computing new type for lambda : " ++ pr_rawconstr rt); + let t' = Pretyping.Default.understand Evd.empty env t in + match n with + | Name id -> + let new_env = Environ.push_rel (n,None,t') env in + let new_b,id_to_exclude = + rebuild_cons new_env + nb_args relname + (args@[mkRVar id])new_crossed_types + (depth + 1 ) b + in + if Idset.mem id id_to_exclude && depth >= nb_args + then + new_b, Idset.remove id (Idset.filter not_free_in_t id_to_exclude) + else + RProd(dummy_loc,n,k,t,new_b),Idset.filter not_free_in_t id_to_exclude + | _ -> anomaly "Should not have an anonymous function here" + (* We have renamed all the anonymous functions during alpha_renaming phase *) + + end + | RLetIn(_,n,t,b) -> + begin + let not_free_in_t id = not (is_free_in id t) in + let t' = Pretyping.Default.understand Evd.empty env t in + let type_t' = Typing.type_of env Evd.empty t' in + let new_env = Environ.push_rel (n,Some 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 Idset.mem id id_to_exclude && depth >= nb_args -> + new_b,Idset.remove id (Idset.filter not_free_in_t id_to_exclude) + | _ -> RLetIn(dummy_loc,n,t,new_b), + Idset.filter not_free_in_t id_to_exclude + end + | RLetTuple(_,nal,(na,rto),t,b) -> + assert (rto=None); + 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' = Pretyping.Default.understand Evd.empty env new_t in + let new_env = Environ.push_rel (na,None,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 Idset.mem id id_to_exclude -> *) +(* new_b,Idset.remove id (Idset.filter not_free_in_t id_to_exclude) *) +(* | _ -> *) + RLetTuple(dummy_loc,nal,(na,None),t,new_b), + Idset.filter not_free_in_t (Idset.union id_to_exclude id_to_exclude') + + end + + | _ -> mkRApp(mkRVar relname,args@[rt]),Idset.empty + + +(* debuging wrapper *) +let rebuild_cons env nb_args relname args crossed_types rt = +(* observennl (str "rebuild_cons : rt := "++ pr_rawconstr rt ++ *) +(* str "nb_args := " ++ str (string_of_int nb_args)); *) + let res = + rebuild_cons env nb_args relname args crossed_types 0 rt + in +(* observe (str " leads to "++ pr_rawconstr (fst res)); *) + res + + +(* naive implementation of parameter detection. + + A parameter is an argument which is only preceded by parameters and whose + calls are all syntaxically equal. + + TODO: Find a valid way to deal with implicit arguments here! +*) +let rec compute_cst_params relnames params = function + | RRef _ | RVar _ | REvar _ | RPatVar _ -> params + | RApp(_,RVar(_,relname'),rtl) when Idset.mem relname' relnames -> + compute_cst_params_from_app [] (params,rtl) + | RApp(_,f,args) -> + List.fold_left (compute_cst_params relnames) params (f::args) + | RLambda(_,_,_,t,b) | RProd(_,_,_,t,b) | RLetIn(_,_,t,b) | RLetTuple(_,_,_,t,b) -> + let t_params = compute_cst_params relnames params t in + compute_cst_params relnames t_params b + | RCases _ -> + params (* If there is still cases at this point they can only be + discriminitation ones *) + | RSort _ -> params + | RHole _ -> params + | RIf _ | RRec _ | RCast _ | RDynamic _ -> + raise (UserError("compute_cst_params", str "Not handled case")) +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,_,is_defined) as param)::params',(RVar(_,id'))::rtl' + when id_ord id id' == 0 && not is_defined -> + compute_cst_params_from_app (param::acc) (params',rtl') + | _ -> List.rev acc + +let compute_params_name relnames (args : (Names.name * Rawterm.rawconstr * bool) 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) + ) + args + in + let l = ref [] in + let _ = + try + list_iter_i + (fun i ((n,nt,is_defined) as param) -> + if array_for_all + (fun l -> + let (n',nt',is_defined') = List.nth l i in + n = n' && Topconstr.eq_rawconstr nt nt' && is_defined = is_defined') + rels_params + then + l := param::!l + ) + rels_params.(0) + with _ -> + () + in + List.rev !l + +let rec rebuild_return_type rt = + match rt with + | Topconstr.CProdN(loc,n,t') -> + Topconstr.CProdN(loc,n,rebuild_return_type t') + | Topconstr.CArrow(loc,t,t') -> + Topconstr.CArrow(loc,t,rebuild_return_type t') + | Topconstr.CLetIn(loc,na,t,t') -> + Topconstr.CLetIn(loc,na,t,rebuild_return_type t') + | _ -> Topconstr.CArrow(dummy_loc,rt,Topconstr.CSort(dummy_loc,RType None)) + + +let do_build_inductive + funnames (funsargs: (Names.name * rawconstr * bool) list list) + returned_types + (rtl:rawconstr list) = + let _time1 = System.get_time () in +(* Pp.msgnl (prlist_with_sep fnl Printer.pr_rawconstr rtl); *) + let funnames_as_set = List.fold_right Idset.add funnames Idset.empty in + let funnames = Array.of_list funnames in + let funsargs = Array.of_list funsargs in + let returned_types = Array.of_list returned_types in + (* alpha_renaming of the body to prevent variable capture during manipulation *) + let rtl_alpha = List.map (function rt -> expand_as (alpha_rt [] rt)) rtl in + let rta = Array.of_list rtl_alpha in + (*i The next call to mk_rel_id is valid since we are constructing the graph + Ensures by: obvious + i*) + let relnames = Array.map mk_rel_id funnames in + let relnames_as_set = Array.fold_right Idset.add relnames Idset.empty in + (* Construction of the pseudo constructors *) + let env = + Array.fold_right + (fun id env -> + Environ.push_named (id,None,Typing.type_of env Evd.empty (Tacinterp.constr_of_id env id)) env + ) + funnames + (Global.env ()) + in + let resa = Array.map (build_entry_lc env funnames_as_set []) rta in + let env_with_graphs = + let rel_arity i funargs = (* Reduilding arities (with parameters) *) + let rel_first_args :(Names.name * Rawterm.rawconstr * bool ) list = + funargs + in + List.fold_right + (fun (n,t,is_defined) acc -> + if is_defined + then + Topconstr.CLetIn(dummy_loc,(dummy_loc, n),Constrextern.extern_rawconstr Idset.empty t, + acc) + else + Topconstr.CProdN + (dummy_loc, + [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_rawconstr Idset.empty t], + 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 mimick a Set Printing All. + Then save the graphs and reset Printing options to their primitive values + *) + let rel_arities = Array.mapi rel_arity funsargs in + Util.array_fold_left2 (fun env rel_name rel_ar -> + Environ.push_named (rel_name,None, Constrintern.interp_constr Evd.empty env rel_ar) env) env relnames rel_arities + in + (* and of the real constructors*) + let constr i res = + List.map + (function result (* (args',concl') *) -> + let rt = compose_raw_context result.context result.value in + let nb_args = List.length funsargs.(i) in + (* with_full_print (fun rt -> Pp.msgnl (str "raw constr " ++ pr_rawconstr rt)) rt; *) + fst ( + rebuild_cons env_with_graphs nb_args relnames.(i) + [] + [] + rt + ) + ) + res.result + in + (* adding names to constructors *) + let next_constructor_id = ref (-1) in + let mk_constructor_id i = + incr next_constructor_id; + (*i The next call to mk_rel_id is valid since we are constructing the graph + Ensures by: obvious + i*) + id_of_string ((string_of_id (mk_rel_id funnames.(i)))^"_"^(string_of_int !next_constructor_id)) + in + let rel_constructors i rt : (identifier*rawconstr) list = + next_constructor_id := (-1); + List.map (fun constr -> (mk_constructor_id i),constr) (constr i rt) + in + let rel_constructors = Array.mapi rel_constructors resa in + (* Computing the set of parameters if asked *) + let rels_params = compute_params_name relnames_as_set funsargs rel_constructors in + let nrel_params = List.length rels_params in + let rel_constructors = (* Taking into account the parameters in constructors *) + Array.map (List.map + (fun (id,rt) -> (id,snd (chop_rprod_n nrel_params rt)))) + rel_constructors + in + let rel_arity i funargs = (* Reduilding arities (with parameters) *) + let rel_first_args :(Names.name * Rawterm.rawconstr * bool ) list = + (snd (list_chop nrel_params funargs)) + in + List.fold_right + (fun (n,t,is_defined) acc -> + if is_defined + then + Topconstr.CLetIn(dummy_loc,(dummy_loc, n),Constrextern.extern_rawconstr Idset.empty t, + acc) + else + Topconstr.CProdN + (dummy_loc, + [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_rawconstr Idset.empty t], + 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 mimick a Set Printing All. + Then save the graphs and reset Printing options to their primitive values + *) + let rel_arities = Array.mapi rel_arity funsargs in + let rel_params = + List.map + (fun (n,t,is_defined) -> + if is_defined + then + Topconstr.LocalRawDef((dummy_loc,n), Constrextern.extern_rawconstr Idset.empty t) + else + Topconstr.LocalRawAssum + ([(dummy_loc,n)], Topconstr.default_binder_kind, Constrextern.extern_rawconstr Idset.empty t) + ) + rels_params + in + let ext_rels_constructors = + Array.map (List.map + (fun (id,t) -> + false,((dummy_loc,id), + Flags.with_option + Flags.raw_print + (Constrextern.extern_rawtype Idset.empty) ((* zeta_normalize *) t) + ) + )) + (rel_constructors) + in + let rel_ind i ext_rel_constructors = + ((dummy_loc,relnames.(i)), + rel_params, + Some rel_arities.(i), + ext_rel_constructors),[] + in + let ext_rel_constructors = (Array.mapi rel_ind ext_rels_constructors) in + let rel_inds = Array.to_list ext_rel_constructors in +(* let _ = *) +(* Pp.msgnl (\* observe *\) ( *) +(* str "Inductive" ++ spc () ++ *) +(* prlist_with_sep *) +(* (fun () -> fnl ()++spc () ++ str "with" ++ spc ()) *) +(* (function ((_,id),_,params,ar,constr) -> *) +(* Ppconstr.pr_id id ++ spc () ++ *) +(* Ppconstr.pr_binders params ++ spc () ++ *) +(* str ":" ++ spc () ++ *) +(* Ppconstr.pr_lconstr_expr ar ++ spc () ++ str ":=" ++ *) +(* prlist_with_sep *) +(* (fun _ -> fnl () ++ spc () ++ str "|" ++ spc ()) *) +(* (function (_,((_,id),t)) -> *) +(* Ppconstr.pr_id id ++ spc () ++ str ":" ++ spc () ++ *) +(* Ppconstr.pr_lconstr_expr t) *) +(* constr *) +(* ) *) +(* rel_inds *) +(* ) *) +(* in *) + let _time2 = System.get_time () in + try + with_full_print (Flags.silently (Command.do_mutual_inductive rel_inds)) true + with + | UserError(s,msg) as e -> + let _time3 = System.get_time () in +(* Pp.msgnl (str "error : "++ str (string_of_float (System.time_difference time2 time3))); *) + let repacked_rel_inds = + List.map (fun ((a , b , c , l),ntn) -> ((false,a) , b, c , Vernacexpr.Inductive_kw, Vernacexpr.Constructors l),ntn ) + rel_inds + in + let msg = + str "while trying to define"++ spc () ++ + Ppvernac.pr_vernac (Vernacexpr.VernacInductive(Decl_kinds.Finite,false,repacked_rel_inds)) + ++ fnl () ++ + msg + in + observe (msg); + raise e + | e -> + let _time3 = System.get_time () in +(* Pp.msgnl (str "error : "++ str (string_of_float (System.time_difference time2 time3))); *) + let repacked_rel_inds = + List.map (fun ((a , b , c , l),ntn) -> ((false,a) , b, c , Vernacexpr.Inductive_kw, Vernacexpr.Constructors l),ntn ) + rel_inds + in + let msg = + str "while trying to define"++ spc () ++ + Ppvernac.pr_vernac (Vernacexpr.VernacInductive(Decl_kinds.Finite,false,repacked_rel_inds)) + ++ fnl () ++ + Cerrors.explain_exn e + in + observe msg; + raise e + + + +let build_inductive funnames funsargs returned_types rtl = + try + do_build_inductive funnames funsargs returned_types rtl + with e -> raise (Building_graph e) + + diff --git a/plugins/funind/rawterm_to_relation.mli b/plugins/funind/rawterm_to_relation.mli new file mode 100644 index 00000000..a314050f --- /dev/null +++ b/plugins/funind/rawterm_to_relation.mli @@ -0,0 +1,16 @@ + + + +(* + [build_inductive parametrize funnames funargs returned_types bodies] + constructs and saves the graphs of the functions [funnames] taking [funargs] as arguments + and returning [returned_types] using bodies [bodies] +*) + +val build_inductive : + Names.identifier list -> (* The list of function name *) + (Names.name*Rawterm.rawconstr*bool) list list -> (* The list of function args *) + Topconstr.constr_expr list -> (* The list of function returned type *) + Rawterm.rawconstr list -> (* the list of body *) + unit + diff --git a/plugins/funind/rawtermops.ml b/plugins/funind/rawtermops.ml new file mode 100644 index 00000000..e31f1452 --- /dev/null +++ b/plugins/funind/rawtermops.ml @@ -0,0 +1,718 @@ +open Pp +open Rawterm +open Util +open Names +(* Ocaml 3.06 Map.S does not handle is_empty *) +let idmap_is_empty m = m = Idmap.empty + +(* + Some basic functions to rebuild rawconstr + In each of them the location is Util.dummy_loc +*) +let mkRRef ref = RRef(dummy_loc,ref) +let mkRVar id = RVar(dummy_loc,id) +let mkRApp(rt,rtl) = RApp(dummy_loc,rt,rtl) +let mkRLambda(n,t,b) = RLambda(dummy_loc,n,Explicit,t,b) +let mkRProd(n,t,b) = RProd(dummy_loc,n,Explicit,t,b) +let mkRLetIn(n,t,b) = RLetIn(dummy_loc,n,t,b) +let mkRCases(rto,l,brl) = RCases(dummy_loc,Term.RegularStyle,rto,l,brl) +let mkRSort s = RSort(dummy_loc,s) +let mkRHole () = RHole(dummy_loc,Evd.BinderType Anonymous) +let mkRCast(b,t) = RCast(dummy_loc,b,CastConv (Term.DEFAULTcast,t)) + +(* + Some basic functions to decompose rawconstrs + These are analogous to the ones constrs +*) +let raw_decompose_prod = + let rec raw_decompose_prod args = function + | RProd(_,n,k,t,b) -> + raw_decompose_prod ((n,t)::args) b + | rt -> args,rt + in + raw_decompose_prod [] + +let raw_decompose_prod_or_letin = + let rec raw_decompose_prod args = function + | RProd(_,n,k,t,b) -> + raw_decompose_prod ((n,None,Some t)::args) b + | RLetIn(_,n,t,b) -> + raw_decompose_prod ((n,Some t,None)::args) b + | rt -> args,rt + in + raw_decompose_prod [] + +let raw_compose_prod = + List.fold_left (fun b (n,t) -> mkRProd(n,t,b)) + +let raw_compose_prod_or_letin = + List.fold_left ( + fun concl decl -> + match decl with + | (n,None,Some t) -> mkRProd(n,t,concl) + | (n,Some bdy,None) -> mkRLetIn(n,bdy,concl) + | _ -> assert false) + +let raw_decompose_prod_n n = + let rec raw_decompose_prod i args c = + if i<=0 then args,c + else + match c with + | RProd(_,n,_,t,b) -> + raw_decompose_prod (i-1) ((n,t)::args) b + | rt -> args,rt + in + raw_decompose_prod n [] + + +let raw_decompose_prod_or_letin_n n = + let rec raw_decompose_prod i args c = + if i<=0 then args,c + else + match c with + | RProd(_,n,_,t,b) -> + raw_decompose_prod (i-1) ((n,None,Some t)::args) b + | RLetIn(_,n,t,b) -> + raw_decompose_prod (i-1) ((n,Some t,None)::args) b + | rt -> args,rt + in + raw_decompose_prod n [] + + +let raw_decompose_app = + let rec decompose_rapp acc rt = +(* msgnl (str "raw_decompose_app on : "++ Printer.pr_rawconstr rt); *) + match rt with + | RApp(_,rt,rtl) -> + decompose_rapp (List.fold_left (fun y x -> x::y) acc rtl) rt + | rt -> rt,List.rev acc + in + decompose_rapp [] + + + + +(* [raw_make_eq t1 t2] build the rawconstr corresponding to [t2 = t1] *) +let raw_make_eq ?(typ= mkRHole ()) t1 t2 = + mkRApp(mkRRef (Lazy.force Coqlib.coq_eq_ref),[typ;t2;t1]) + +(* [raw_make_neq t1 t2] build the rawconstr corresponding to [t1 <> t2] *) +let raw_make_neq t1 t2 = + mkRApp(mkRRef (Lazy.force Coqlib.coq_not_ref),[raw_make_eq t1 t2]) + +(* [raw_make_or P1 P2] build the rawconstr corresponding to [P1 \/ P2] *) +let raw_make_or t1 t2 = mkRApp (mkRRef(Lazy.force Coqlib.coq_or_ref),[t1;t2]) + +(* [raw_make_or_list [P1;...;Pn]] build the rawconstr corresponding + to [P1 \/ ( .... \/ Pn)] +*) +let rec raw_make_or_list = function + | [] -> raise (Invalid_argument "mk_or") + | [e] -> e + | e::l -> raw_make_or e (raw_make_or_list l) + + +let remove_name_from_mapping mapping na = + match na with + | Anonymous -> mapping + | Name id -> Idmap.remove id mapping + +let change_vars = + let rec change_vars mapping rt = + match rt with + | RRef _ -> rt + | RVar(loc,id) -> + let new_id = + try + Idmap.find id mapping + with Not_found -> id + in + RVar(loc,new_id) + | REvar _ -> rt + | RPatVar _ -> rt + | RApp(loc,rt',rtl) -> + RApp(loc, + change_vars mapping rt', + List.map (change_vars mapping) rtl + ) + | RLambda(loc,name,k,t,b) -> + RLambda(loc, + name, + k, + change_vars mapping t, + change_vars (remove_name_from_mapping mapping name) b + ) + | RProd(loc,name,k,t,b) -> + RProd(loc, + name, + k, + change_vars mapping t, + change_vars (remove_name_from_mapping mapping name) b + ) + | RLetIn(loc,name,def,b) -> + RLetIn(loc, + name, + change_vars mapping def, + change_vars (remove_name_from_mapping mapping name) b + ) + | RLetTuple(loc,nal,(na,rto),b,e) -> + let new_mapping = List.fold_left remove_name_from_mapping mapping nal in + RLetTuple(loc, + nal, + (na, Option.map (change_vars mapping) rto), + change_vars mapping b, + change_vars new_mapping e + ) + | RCases(loc,sty,infos,el,brl) -> + RCases(loc,sty, + infos, + List.map (fun (e,x) -> (change_vars mapping e,x)) el, + List.map (change_vars_br mapping) brl + ) + | RIf(loc,b,(na,e_option),lhs,rhs) -> + RIf(loc, + change_vars mapping b, + (na,Option.map (change_vars mapping) e_option), + change_vars mapping lhs, + change_vars mapping rhs + ) + | RRec _ -> error "Local (co)fixes are not supported" + | RSort _ -> rt + | RHole _ -> rt + | RCast(loc,b,CastConv (k,t)) -> + RCast(loc,change_vars mapping b, CastConv (k,change_vars mapping t)) + | RCast(loc,b,CastCoerce) -> + RCast(loc,change_vars mapping b,CastCoerce) + | RDynamic _ -> error "Not handled RDynamic" + and change_vars_br mapping ((loc,idl,patl,res) as br) = + let new_mapping = List.fold_right Idmap.remove idl mapping in + if idmap_is_empty new_mapping + then br + else (loc,idl,patl,change_vars new_mapping res) + in + change_vars + + + +let rec alpha_pat excluded pat = + match pat with + | PatVar(loc,Anonymous) -> + let new_id = Indfun_common.fresh_id excluded "_x" in + PatVar(loc,Name new_id),(new_id::excluded),Idmap.empty + | PatVar(loc,Name id) -> + if List.mem id excluded + then + let new_id = Namegen.next_ident_away id excluded in + PatVar(loc,Name new_id),(new_id::excluded), + (Idmap.add id new_id Idmap.empty) + else pat,excluded,Idmap.empty + | PatCstr(loc,constr,patl,na) -> + let new_na,new_excluded,map = + match na with + | Name id when List.mem id excluded -> + let new_id = Namegen.next_ident_away id excluded in + Name new_id,new_id::excluded, Idmap.add id new_id Idmap.empty + | _ -> na,excluded,Idmap.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,Idmap.fold Idmap.add new_map map) + ) + ([],new_excluded,map) + patl + in + PatCstr(loc,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,(Idmap.fold Idmap.add new_map map) + ) + ([],excluded,Idmap.empty) + patl + in + (List.rev patl,new_excluded,map) + + + + +let raw_get_pattern_id pat acc = + let rec get_pattern_id pat = + match pat with + | PatVar(loc,Anonymous) -> assert false + | PatVar(loc,Name id) -> + [id] + | PatCstr(loc,constr,patternl,_) -> + List.fold_right + (fun pat idl -> + let idl' = get_pattern_id pat in + idl'@idl + ) + patternl + [] + in + (get_pattern_id pat)@acc + +let get_pattern_id pat = raw_get_pattern_id pat [] + +let rec alpha_rt excluded rt = + let new_rt = + match rt with + | RRef _ | RVar _ | REvar _ | RPatVar _ -> rt + | RLambda(loc,Anonymous,k,t,b) -> + let new_id = Namegen.next_ident_away (id_of_string "_x") 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 + RLambda(loc,Name new_id,k,new_t,new_b) + | RProd(loc,Anonymous,k,t,b) -> + let new_t = alpha_rt excluded t in + let new_b = alpha_rt excluded b in + RProd(loc,Anonymous,k,new_t,new_b) + | RLetIn(loc,Anonymous,t,b) -> + let new_t = alpha_rt excluded t in + let new_b = alpha_rt excluded b in + RLetIn(loc,Anonymous,new_t,new_b) + | RLambda(loc,Name id,k,t,b) -> + let new_id = Namegen.next_ident_away id excluded in + let t,b = + if new_id = id + then t,b + else + let replace = change_vars (Idmap.add id new_id Idmap.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 + RLambda(loc,Name new_id,k,new_t,new_b) + | RProd(loc,Name id,k,t,b) -> + let new_id = Namegen.next_ident_away id excluded in + let new_excluded = new_id::excluded in + let t,b = + if new_id = id + then t,b + else + let replace = change_vars (Idmap.add id new_id Idmap.empty) in + (t,replace b) + in + let new_t = alpha_rt new_excluded t in + let new_b = alpha_rt new_excluded b in + RProd(loc,Name new_id,k,new_t,new_b) + | RLetIn(loc,Name id,t,b) -> + let new_id = Namegen.next_ident_away id excluded in + let t,b = + if new_id = id + then t,b + else + let replace = change_vars (Idmap.add id new_id Idmap.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 + RLetIn(loc,Name new_id,new_t,new_b) + + + | RLetTuple(loc,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 excluded in + if new_id = id + then + na::nal,id::excluded,mapping + else + (Name new_id)::nal,id::excluded,(Idmap.add id new_id mapping) + ) + ([],excluded,Idmap.empty) + nal + in + let new_nal = List.rev rev_new_nal in + let new_rto,new_t,new_b = + if idmap_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 + RLetTuple(loc,new_nal,(na,new_rto),new_t,new_b) + | RCases(loc,sty,infos,el,brl) -> + let new_el = + List.map (function (rt,i) -> alpha_rt excluded rt, i) el + in + RCases(loc,sty,infos,new_el,List.map (alpha_br excluded) brl) + | RIf(loc,b,(na,e_o),lhs,rhs) -> + RIf(loc,alpha_rt excluded b, + (na,Option.map (alpha_rt excluded) e_o), + alpha_rt excluded lhs, + alpha_rt excluded rhs + ) + | RRec _ -> error "Not handled RRec" + | RSort _ -> rt + | RHole _ -> rt + | RCast (loc,b,CastConv (k,t)) -> + RCast(loc,alpha_rt excluded b,CastConv(k,alpha_rt excluded t)) + | RCast (loc,b,CastCoerce) -> + RCast(loc,alpha_rt excluded b,CastCoerce) + | RDynamic _ -> error "Not handled RDynamic" + | RApp(loc,f,args) -> + RApp(loc, + alpha_rt excluded f, + List.map (alpha_rt excluded) args + ) + in + new_rt + +and alpha_br excluded (loc,ids,patl,res) = + let new_patl,new_excluded,mapping = alpha_patl excluded patl in + let new_ids = List.fold_right raw_get_pattern_id new_patl [] in + let new_excluded = new_ids@excluded in + let renamed_res = change_vars mapping res in + let new_res = alpha_rt new_excluded renamed_res in + (loc,new_ids,new_patl,new_res) + +(* + [is_free_in id rt] checks if [id] is a free variable in [rt] +*) +let is_free_in id = + let rec is_free_in = function + | RRef _ -> false + | RVar(_,id') -> id_ord id' id == 0 + | REvar _ -> false + | RPatVar _ -> false + | RApp(_,rt,rtl) -> List.exists is_free_in (rt::rtl) + | RLambda(_,n,_,t,b) | RProd(_,n,_,t,b) | RLetIn(_,n,t,b) -> + let check_in_b = + match n with + | Name id' -> id_ord id' id <> 0 + | _ -> true + in + is_free_in t || (check_in_b && is_free_in b) + | RCases(_,_,_,el,brl) -> + (List.exists (fun (e,_) -> is_free_in e) el) || + List.exists is_free_in_br brl + | RLetTuple(_,nal,_,b,t) -> + let check_in_nal = + not (List.exists (function Name id' -> id'= id | _ -> false) nal) + in + is_free_in t || (check_in_nal && is_free_in b) + + | RIf(_,cond,_,br1,br2) -> + is_free_in cond || is_free_in br1 || is_free_in br2 + | RRec _ -> raise (UserError("",str "Not handled RRec")) + | RSort _ -> false + | RHole _ -> false + | RCast (_,b,CastConv (_,t)) -> is_free_in b || is_free_in t + | RCast (_,b,CastCoerce) -> is_free_in b + | RDynamic _ -> raise (UserError("",str "Not handled RDynamic")) + and is_free_in_br (_,ids,_,rt) = + (not (List.mem id ids)) && is_free_in rt + in + is_free_in + + + +let rec pattern_to_term = function + | PatVar(loc,Anonymous) -> assert false + | PatVar(loc,Name id) -> + mkRVar id + | PatCstr(loc,constr,patternl,_) -> + let cst_narg = + Inductiveops.mis_constructor_nargs_env + (Global.env ()) + constr + in + let implicit_args = + Array.to_list + (Array.init + (cst_narg - List.length patternl) + (fun _ -> mkRHole ()) + ) + in + let patl_as_term = + List.map pattern_to_term patternl + in + mkRApp(mkRRef(Libnames.ConstructRef constr), + implicit_args@patl_as_term + ) + + + +let replace_var_by_term x_id term = + let rec replace_var_by_pattern rt = + match rt with + | RRef _ -> rt + | RVar(_,id) when id_ord id x_id == 0 -> term + | RVar _ -> rt + | REvar _ -> rt + | RPatVar _ -> rt + | RApp(loc,rt',rtl) -> + RApp(loc, + replace_var_by_pattern rt', + List.map replace_var_by_pattern rtl + ) + | RLambda(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt + | RLambda(loc,name,k,t,b) -> + RLambda(loc, + name, + k, + replace_var_by_pattern t, + replace_var_by_pattern b + ) + | RProd(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt + | RProd(loc,name,k,t,b) -> + RProd(loc, + name, + k, + replace_var_by_pattern t, + replace_var_by_pattern b + ) + | RLetIn(_,Name id,_,_) when id_ord id x_id == 0 -> rt + | RLetIn(loc,name,def,b) -> + RLetIn(loc, + name, + replace_var_by_pattern def, + replace_var_by_pattern b + ) + | RLetTuple(_,nal,_,_,_) + when List.exists (function Name id -> id = x_id | _ -> false) nal -> + rt + | RLetTuple(loc,nal,(na,rto),def,b) -> + RLetTuple(loc, + nal, + (na,Option.map replace_var_by_pattern rto), + replace_var_by_pattern def, + replace_var_by_pattern b + ) + | RCases(loc,sty,infos,el,brl) -> + RCases(loc,sty, + infos, + List.map (fun (e,x) -> (replace_var_by_pattern e,x)) el, + List.map replace_var_by_pattern_br brl + ) + | RIf(loc,b,(na,e_option),lhs,rhs) -> + RIf(loc, replace_var_by_pattern b, + (na,Option.map replace_var_by_pattern e_option), + replace_var_by_pattern lhs, + replace_var_by_pattern rhs + ) + | RRec _ -> raise (UserError("",str "Not handled RRec")) + | RSort _ -> rt + | RHole _ -> rt + | RCast(loc,b,CastConv(k,t)) -> + RCast(loc,replace_var_by_pattern b,CastConv(k,replace_var_by_pattern t)) + | RCast(loc,b,CastCoerce) -> + RCast(loc,replace_var_by_pattern b,CastCoerce) + | RDynamic _ -> raise (UserError("",str "Not handled RDynamic")) + and replace_var_by_pattern_br ((loc,idl,patl,res) as br) = + if List.exists (fun id -> id_ord id x_id == 0) idl + then br + else (loc,idl,patl,replace_var_by_pattern res) + in + replace_var_by_pattern + + + + +(* checking unifiability of patterns *) +exception NotUnifiable + +let rec are_unifiable_aux = function + | [] -> () + | eq::eqs -> + match eq with + | PatVar _,_ | _,PatVar _ -> are_unifiable_aux eqs + | PatCstr(_,constructor1,cpl1,_),PatCstr(_,constructor2,cpl2,_) -> + if constructor2 <> constructor1 + then raise NotUnifiable + else + let eqs' = + try ((List.combine cpl1 cpl2)@eqs) + with _ -> anomaly "are_unifiable_aux" + in + are_unifiable_aux eqs' + +let are_unifiable pat1 pat2 = + try + are_unifiable_aux [pat1,pat2]; + true + with NotUnifiable -> false + + +let rec eq_cases_pattern_aux = function + | [] -> () + | eq::eqs -> + match eq with + | PatVar _,PatVar _ -> eq_cases_pattern_aux eqs + | PatCstr(_,constructor1,cpl1,_),PatCstr(_,constructor2,cpl2,_) -> + if constructor2 <> constructor1 + then raise NotUnifiable + else + let eqs' = + try ((List.combine cpl1 cpl2)@eqs) + with _ -> anomaly "eq_cases_pattern_aux" + in + eq_cases_pattern_aux eqs' + | _ -> raise NotUnifiable + +let eq_cases_pattern pat1 pat2 = + try + eq_cases_pattern_aux [pat1,pat2]; + true + with NotUnifiable -> false + + + +let ids_of_pat = + let rec ids_of_pat ids = function + | PatVar(_,Anonymous) -> ids + | PatVar(_,Name id) -> Idset.add id ids + | PatCstr(_,_,patl,_) -> List.fold_left ids_of_pat ids patl + in + ids_of_pat Idset.empty + +let id_of_name = function + | Names.Anonymous -> id_of_string "x" + | Names.Name x -> x + +(* TODO: finish Rec caes *) +let ids_of_rawterm c = + let rec ids_of_rawterm acc c = + let idof = id_of_name in + match c with + | RVar (_,id) -> id::acc + | RApp (loc,g,args) -> + ids_of_rawterm [] g @ List.flatten (List.map (ids_of_rawterm []) args) @ acc + | RLambda (loc,na,k,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc + | RProd (loc,na,k,ty,c) -> idof na :: ids_of_rawterm [] ty @ ids_of_rawterm [] c @ acc + | RLetIn (loc,na,b,c) -> idof na :: ids_of_rawterm [] b @ ids_of_rawterm [] c @ acc + | RCast (loc,c,CastConv(k,t)) -> ids_of_rawterm [] c @ ids_of_rawterm [] t @ acc + | RCast (loc,c,CastCoerce) -> ids_of_rawterm [] c @ acc + | RIf (loc,c,(na,po),b1,b2) -> ids_of_rawterm [] c @ ids_of_rawterm [] b1 @ ids_of_rawterm [] b2 @ acc + | RLetTuple (_,nal,(na,po),b,c) -> + List.map idof nal @ ids_of_rawterm [] b @ ids_of_rawterm [] c @ acc + | RCases (loc,sty,rtntypopt,tml,brchl) -> + List.flatten (List.map (fun (_,idl,patl,c) -> idl @ ids_of_rawterm [] c) brchl) + | RRec _ -> failwith "Fix inside a constructor branch" + | (RSort _ | RHole _ | RRef _ | REvar _ | RPatVar _ | RDynamic _) -> [] + in + (* build the set *) + List.fold_left (fun acc x -> Idset.add x acc) Idset.empty (ids_of_rawterm [] c) + + + + + +let zeta_normalize = + let rec zeta_normalize_term rt = + match rt with + | RRef _ -> rt + | RVar _ -> rt + | REvar _ -> rt + | RPatVar _ -> rt + | RApp(loc,rt',rtl) -> + RApp(loc, + zeta_normalize_term rt', + List.map zeta_normalize_term rtl + ) + | RLambda(loc,name,k,t,b) -> + RLambda(loc, + name, + k, + zeta_normalize_term t, + zeta_normalize_term b + ) + | RProd(loc,name,k,t,b) -> + RProd(loc, + name, + k, + zeta_normalize_term t, + zeta_normalize_term b + ) + | RLetIn(_,Name id,def,b) -> + zeta_normalize_term (replace_var_by_term id def b) + | RLetIn(loc,Anonymous,def,b) -> zeta_normalize_term b + | RLetTuple(loc,nal,(na,rto),def,b) -> + RLetTuple(loc, + nal, + (na,Option.map zeta_normalize_term rto), + zeta_normalize_term def, + zeta_normalize_term b + ) + | RCases(loc,sty,infos,el,brl) -> + RCases(loc,sty, + infos, + List.map (fun (e,x) -> (zeta_normalize_term e,x)) el, + List.map zeta_normalize_br brl + ) + | RIf(loc,b,(na,e_option),lhs,rhs) -> + RIf(loc, zeta_normalize_term b, + (na,Option.map zeta_normalize_term e_option), + zeta_normalize_term lhs, + zeta_normalize_term rhs + ) + | RRec _ -> raise (UserError("",str "Not handled RRec")) + | RSort _ -> rt + | RHole _ -> rt + | RCast(loc,b,CastConv(k,t)) -> + RCast(loc,zeta_normalize_term b,CastConv(k,zeta_normalize_term t)) + | RCast(loc,b,CastCoerce) -> + RCast(loc,zeta_normalize_term b,CastCoerce) + | RDynamic _ -> raise (UserError("",str "Not handled RDynamic")) + and zeta_normalize_br (loc,idl,patl,res) = + (loc,idl,patl,zeta_normalize_term res) + in + zeta_normalize_term + + + + +let expand_as = + + let rec add_as map pat = + match pat with + | PatVar _ -> map + | PatCstr(_,_,patl,Name id) -> + Idmap.add id (pattern_to_term pat) (List.fold_left add_as map patl) + | PatCstr(_,_,patl,_) -> List.fold_left add_as map patl + in + let rec expand_as map rt = + match rt with + | RRef _ | REvar _ | RPatVar _ | RSort _ | RHole _ -> rt + | RVar(_,id) -> + begin + try + Idmap.find id map + with Not_found -> rt + end + | RApp(loc,f,args) -> RApp(loc,expand_as map f,List.map (expand_as map) args) + | RLambda(loc,na,k,t,b) -> RLambda(loc,na,k,expand_as map t, expand_as map b) + | RProd(loc,na,k,t,b) -> RProd(loc,na,k,expand_as map t, expand_as map b) + | RLetIn(loc,na,v,b) -> RLetIn(loc,na, expand_as map v,expand_as map b) + | RLetTuple(loc,nal,(na,po),v,b) -> + RLetTuple(loc,nal,(na,Option.map (expand_as map) po), + expand_as map v, expand_as map b) + | RIf(loc,e,(na,po),br1,br2) -> + RIf(loc,expand_as map e,(na,Option.map (expand_as map) po), + expand_as map br1, expand_as map br2) + | RRec _ -> error "Not handled RRec" + | RDynamic _ -> error "Not handled RDynamic" + | RCast(loc,b,CastConv(kind,t)) -> RCast(loc,expand_as map b,CastConv(kind,expand_as map t)) + | RCast(loc,b,CastCoerce) -> RCast(loc,expand_as map b,CastCoerce) + | RCases(loc,sty,po,el,brl) -> + RCases(loc, 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 (loc,idl,cpl,rt) = + (loc,idl,cpl, expand_as (List.fold_left add_as map cpl) rt) + in + expand_as Idmap.empty diff --git a/plugins/funind/rawtermops.mli b/plugins/funind/rawtermops.mli new file mode 100644 index 00000000..455e7c89 --- /dev/null +++ b/plugins/funind/rawtermops.mli @@ -0,0 +1,126 @@ +open Rawterm + +(* Ocaml 3.06 Map.S does not handle is_empty *) +val idmap_is_empty : 'a Names.Idmap.t -> bool + + +(* [get_pattern_id pat] returns a list of all the variable appearing in [pat] *) +val get_pattern_id : cases_pattern -> Names.identifier list + +(* [pattern_to_term pat] returns a rawconstr corresponding to [pat]. + [pat] must not contain occurences of anonymous pattern +*) +val pattern_to_term : cases_pattern -> rawconstr + +(* + Some basic functions to rebuild rawconstr + In each of them the location is Util.dummy_loc +*) +val mkRRef : Libnames.global_reference -> rawconstr +val mkRVar : Names.identifier -> rawconstr +val mkRApp : rawconstr*(rawconstr list) -> rawconstr +val mkRLambda : Names.name*rawconstr*rawconstr -> rawconstr +val mkRProd : Names.name*rawconstr*rawconstr -> rawconstr +val mkRLetIn : Names.name*rawconstr*rawconstr -> rawconstr +val mkRCases : rawconstr option * tomatch_tuples * cases_clauses -> rawconstr +val mkRSort : rawsort -> rawconstr +val mkRHole : unit -> rawconstr (* we only build Evd.BinderType Anonymous holes *) +val mkRCast : rawconstr* rawconstr -> rawconstr +(* + Some basic functions to decompose rawconstrs + These are analogous to the ones constrs +*) +val raw_decompose_prod : rawconstr -> (Names.name*rawconstr) list * rawconstr +val raw_decompose_prod_or_letin : + rawconstr -> (Names.name*rawconstr option*rawconstr option) list * rawconstr +val raw_decompose_prod_n : int -> rawconstr -> (Names.name*rawconstr) list * rawconstr +val raw_decompose_prod_or_letin_n : int -> rawconstr -> + (Names.name*rawconstr option*rawconstr option) list * rawconstr +val raw_compose_prod : rawconstr -> (Names.name*rawconstr) list -> rawconstr +val raw_compose_prod_or_letin: rawconstr -> + (Names.name*rawconstr option*rawconstr option) list -> rawconstr +val raw_decompose_app : rawconstr -> rawconstr*(rawconstr list) + + +(* [raw_make_eq t1 t2] build the rawconstr corresponding to [t2 = t1] *) +val raw_make_eq : ?typ:rawconstr -> rawconstr -> rawconstr -> rawconstr +(* [raw_make_neq t1 t2] build the rawconstr corresponding to [t1 <> t2] *) +val raw_make_neq : rawconstr -> rawconstr -> rawconstr +(* [raw_make_or P1 P2] build the rawconstr corresponding to [P1 \/ P2] *) +val raw_make_or : rawconstr -> rawconstr -> rawconstr + +(* [raw_make_or_list [P1;...;Pn]] build the rawconstr corresponding + to [P1 \/ ( .... \/ Pn)] +*) +val raw_make_or_list : rawconstr list -> rawconstr + + +(* alpha_conversion functions *) + + + +(* Replace the var mapped in the rawconstr/context *) +val change_vars : Names.identifier Names.Idmap.t -> rawconstr -> rawconstr + + + +(* [alpha_pat avoid pat] rename all the variables present in [pat] s.t. + the result does not share variables with [avoid]. This function create + a fresh variable for each occurence of the anonymous pattern. + + Also returns a mapping from old variables to new ones and the concatenation of + [avoid] with the variables appearing in the result. +*) + val alpha_pat : + Names.Idmap.key list -> + Rawterm.cases_pattern -> + Rawterm.cases_pattern * Names.Idmap.key list * + Names.identifier Names.Idmap.t + +(* [alpha_rt avoid rt] alpha convert [rt] s.t. the result repects barendregt + conventions and does not share bound variables with avoid +*) +val alpha_rt : Names.identifier list -> rawconstr -> rawconstr + +(* same as alpha_rt but for case branches *) +val alpha_br : Names.identifier list -> + Util.loc * Names.identifier list * Rawterm.cases_pattern list * + Rawterm.rawconstr -> + Util.loc * Names.identifier list * Rawterm.cases_pattern list * + Rawterm.rawconstr + + +(* Reduction function *) +val replace_var_by_term : + Names.identifier -> + Rawterm.rawconstr -> Rawterm.rawconstr -> Rawterm.rawconstr + + + +(* + [is_free_in id rt] checks if [id] is a free variable in [rt] +*) +val is_free_in : Names.identifier -> rawconstr -> bool + + +val are_unifiable : cases_pattern -> cases_pattern -> bool +val eq_cases_pattern : cases_pattern -> cases_pattern -> bool + + + +(* + ids_of_pat : cases_pattern -> Idset.t + returns the set of variables appearing in a pattern +*) +val ids_of_pat : cases_pattern -> Names.Idset.t + +(* TODO: finish this function (Fix not treated) *) +val ids_of_rawterm: rawconstr -> Names.Idset.t + +(* + removing let_in construction in a rawterm +*) +val zeta_normalize : Rawterm.rawconstr -> Rawterm.rawconstr + + +val expand_as : rawconstr -> rawconstr diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml new file mode 100644 index 00000000..3b0b8628 --- /dev/null +++ b/plugins/funind/recdef.ml @@ -0,0 +1,1473 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* next_global_ident_away id (acc@ids)::acc) + idl + [] + +let pf_get_new_id id g = + List.hd (pf_get_new_ids [id] g) + +let h_intros l = + tclMAP h_intro l + +let debug_queue = Queue.create () + + +let rec print_debug_queue e = + let lmsg,goal = Queue.pop debug_queue in + if Queue.is_empty debug_queue + then + msgnl (lmsg ++ (str " raised exception " ++ Cerrors.explain_exn e) ++ str " on goal " ++ goal) + else + begin + print_debug_queue e; + msgnl (str " from " ++ lmsg ++ str " on goal " ++ goal); + end + + +let do_observe_tac s tac g = + let goal = Printer.pr_goal (sig_it g) in + let lmsg = (str "recdef ") ++ (str s) in + Queue.add (lmsg,goal) debug_queue; + try + let v = tac g in + ignore(Queue.pop debug_queue); + v + with e -> + if not (Queue.is_empty debug_queue) + then + print_debug_queue e; + raise e + +(*let do_observe_tac s tac g = + let goal = begin (Printer.pr_goal (sig_it g)) end in + try let v = tac g in msgnl (goal ++ fnl () ++ (str "recdef ") ++ + (str s)++(str " ")++(str "finished")); v + with e -> + msgnl (str "observation "++str s++str " raised exception " ++ + Cerrors.explain_exn e ++ str " on goal " ++ goal ); + raise e;; +*) + +let observe_tac s tac g = + if Tacinterp.get_debug () <> Tactic_debug.DebugOff + then do_observe_tac s tac g + else tac g + +let hyp_ids = List.map id_of_string + ["x";"v";"k";"def";"p";"h";"n";"h'"; "anonymous"; "teq"; "rec_res"; + "hspec";"heq"; "hrec"; "hex"; "teq"; "pmax";"hle"];; + +let rec nthtl = function + l, 0 -> l | _::tl, n -> nthtl (tl, n-1) | [], _ -> [];; + +let hyp_id n l = List.nth l n;; + +let (x_id:identifier) = hyp_id 0 hyp_ids;; +let (v_id:identifier) = hyp_id 1 hyp_ids;; +let (k_id:identifier) = hyp_id 2 hyp_ids;; +let (def_id:identifier) = hyp_id 3 hyp_ids;; +let (p_id:identifier) = hyp_id 4 hyp_ids;; +let (h_id:identifier) = hyp_id 5 hyp_ids;; +let (n_id:identifier) = hyp_id 6 hyp_ids;; +let (h'_id:identifier) = hyp_id 7 hyp_ids;; +let (ano_id:identifier) = hyp_id 8 hyp_ids;; +let (rec_res_id:identifier) = hyp_id 10 hyp_ids;; +let (hspec_id:identifier) = hyp_id 11 hyp_ids;; +let (heq_id:identifier) = hyp_id 12 hyp_ids;; +let (hrec_id:identifier) = hyp_id 13 hyp_ids;; +let (hex_id:identifier) = hyp_id 14 hyp_ids;; +let (teq_id:identifier) = hyp_id 15 hyp_ids;; +let (pmax_id:identifier) = hyp_id 16 hyp_ids;; +let (hle_id:identifier) = hyp_id 17 hyp_ids;; + +let message s = if Flags.is_verbose () then msgnl(str s);; + +let def_of_const t = + match (kind_of_term t) with + Const sp -> + (try (match (Global.lookup_constant sp) with + {const_body=Some c} -> Declarations.force c + |_ -> assert false) + with _ -> + anomaly ("Cannot find definition of constant "^ + (string_of_id (id_of_label (con_label sp)))) + ) + |_ -> assert false + +let type_of_const t = + match (kind_of_term t) with + Const sp -> Typeops.type_of_constant (Global.env()) sp + |_ -> assert false + +let arg_type t = + match kind_of_term (def_of_const t) with + Lambda(a,b,c) -> b + | _ -> assert false;; + +let evaluable_of_global_reference r = + match r with + ConstRef sp -> EvalConstRef sp + | VarRef id -> EvalVarRef id + | _ -> assert false;; + + +let rank_for_arg_list h = + let predicate a b = + try List.for_all2 eq_constr a b with + Invalid_argument _ -> false in + let rec rank_aux i = function + | [] -> None + | x::tl -> if predicate h x then Some i else rank_aux (i+1) tl in + rank_aux 0;; + +let rec (find_call_occs : int -> constr -> constr -> + (constr list -> constr) * constr list list) = + fun nb_lam f expr -> + match (kind_of_term expr) with + App (g, args) when g = f -> + (fun l -> List.hd l), [Array.to_list args] + | App (g, args) -> + let (largs: constr list) = Array.to_list args in + let rec find_aux = function + [] -> (fun x -> []), [] + | a::upper_tl -> + (match find_aux upper_tl with + (cf, ((arg1::args) as args_for_upper_tl)) -> + (match find_call_occs nb_lam f a with + cf2, (_ :: _ as other_args) -> + let rec avoid_duplicates args = + match args with + | [] -> (fun _ -> []), [] + | h::tl -> + let recomb_tl, args_for_tl = + avoid_duplicates tl in + match rank_for_arg_list h args_for_upper_tl with + | None -> + (fun l -> List.hd l::recomb_tl(List.tl l)), + h::args_for_tl + | Some i -> + (fun l -> List.nth l (i+List.length args_for_tl):: + recomb_tl l), + args_for_tl + in + let recombine, other_args' = + avoid_duplicates other_args in + let len1 = List.length other_args' in + (fun l -> cf2 (recombine l)::cf(nthtl(l,len1))), + other_args'@args_for_upper_tl + | _, [] -> (fun x -> a::cf x), args_for_upper_tl) + | _, [] -> + (match find_call_occs nb_lam f a with + cf, (arg1::args) -> (fun l -> cf l::upper_tl), (arg1::args) + | _, [] -> (fun x -> a::upper_tl), [])) in + begin + match (find_aux largs) with + cf, [] -> (fun l -> mkApp(g, args)), [] + | cf, args -> + (fun l -> mkApp (g, Array.of_list (cf l))), args + end + | Rel(v) -> if v > nb_lam then error "find_call_occs : Rel" else ((fun l -> expr),[]) + | Var(id) -> (fun l -> expr), [] + | Meta(_) -> error "find_call_occs : Meta" + | Evar(_) -> error "find_call_occs : Evar" + | Sort(_) -> (fun l -> expr), [] + | Cast(b,_,_) -> find_call_occs nb_lam f b + | Prod(_,_,_) -> error "find_call_occs : Prod" + | Lambda(na,t,b) -> + begin + match find_call_occs (succ nb_lam) f b with + | _, [] -> (* Lambda are authorized as long as they do not contain + recursives calls *) + (fun l -> expr),[] + | _ -> error "find_call_occs : Lambda" + end + | LetIn(na,v,t,b) -> + begin + match find_call_occs nb_lam f v, find_call_occs (succ nb_lam) f b with + | (_,[]),(_,[]) -> + ((fun l -> expr), []) + | (_,[]),(cf,(_::_ as l)) -> + ((fun l -> mkLetIn(na,v,t,cf l)),l) + | (cf,(_::_ as l)),(_,[]) -> + ((fun l -> mkLetIn(na,cf l,t,b)), l) + | _ -> error "find_call_occs : LetIn" + end + | Const(_) -> (fun l -> expr), [] + | Ind(_) -> (fun l -> expr), [] + | Construct (_, _) -> (fun l -> expr), [] + | Case(i,t,a,r) -> + (match find_call_occs nb_lam f a with + cf, (arg1::args) -> (fun l -> mkCase(i, t, (cf l), r)),(arg1::args) + | _ -> (fun l -> expr),[]) + | Fix(_) -> error "find_call_occs : Fix" + | CoFix(_) -> error "find_call_occs : CoFix";; + +let coq_constant s = + Coqlib.gen_constant_in_modules "RecursiveDefinition" + (Coqlib.init_modules @ Coqlib.arith_modules) s;; + +let coq_base_constant s = + Coqlib.gen_constant_in_modules "RecursiveDefinition" + (Coqlib.init_modules @ [["Coq";"Arith";"Le"];["Coq";"Arith";"Lt"]]) s;; + +let constant sl s = + constr_of_global + (locate (make_qualid(Names.make_dirpath + (List.map id_of_string (List.rev sl))) + (id_of_string s)));; + +let find_reference sl s = + (locate (make_qualid(Names.make_dirpath + (List.map id_of_string (List.rev sl))) + (id_of_string s)));; + +let delayed_force f = f () + +let le_lt_SS = function () -> (constant ["Recdef"] "le_lt_SS") +let le_lt_n_Sm = function () -> (coq_base_constant "le_lt_n_Sm") + +let le_trans = function () -> (coq_base_constant "le_trans") +let le_lt_trans = function () -> (coq_base_constant "le_lt_trans") +let lt_S_n = function () -> (coq_base_constant "lt_S_n") +let le_n = function () -> (coq_base_constant "le_n") +let refl_equal = function () -> (coq_base_constant "eq_refl") +let eq = function () -> (coq_base_constant "eq") +let ex = function () -> (coq_base_constant "ex") +let coq_sig_ref = function () -> (find_reference ["Coq";"Init";"Specif"] "sig") +let coq_sig = function () -> (coq_base_constant "sig") +let coq_O = function () -> (coq_base_constant "O") +let coq_S = function () -> (coq_base_constant "S") + +let gt_antirefl = function () -> (coq_constant "gt_irrefl") +let lt_n_O = function () -> (coq_base_constant "lt_n_O") +let lt_n_Sn = function () -> (coq_base_constant "lt_n_Sn") + +let f_equal = function () -> (coq_constant "f_equal") +let well_founded_induction = function () -> (coq_constant "well_founded_induction") +let well_founded = function () -> (coq_constant "well_founded") +let acc_rel = function () -> (coq_constant "Acc") +let acc_inv_id = function () -> (coq_constant "Acc_inv") +let well_founded_ltof = function () -> (Coqlib.coq_constant "" ["Arith";"Wf_nat"] "well_founded_ltof") +let iter_ref = function () -> (try find_reference ["Recdef"] "iter" with Not_found -> error "module Recdef not loaded") +let max_ref = function () -> (find_reference ["Recdef"] "max") +let iter = function () -> (constr_of_global (delayed_force iter_ref)) +let max_constr = function () -> (constr_of_global (delayed_force max_ref)) + +let ltof_ref = function () -> (find_reference ["Coq";"Arith";"Wf_nat"] "ltof") +let coq_conj = function () -> find_reference ["Coq";"Init";"Logic"] "conj" + +(* These are specific to experiments in nat with lt as well_founded_relation, *) +(* but this should be made more general. *) +let nat = function () -> (coq_base_constant "nat") +let lt = function () -> (coq_base_constant "lt") + +(* This is simply an implementation of the case_eq tactic. this code + should be replaced with the tactic defined in Ltac in Init/Tactics.v *) +let mkCaseEq a : tactic = + (fun g -> + let type_of_a = pf_type_of g a in + tclTHENLIST + [h_generalize [mkApp(delayed_force refl_equal, [| type_of_a; a|])]; + (fun g2 -> + change_in_concl None + (pattern_occs [((false,[1]), a)] (pf_env g2) Evd.empty (pf_concl g2)) + g2); + simplest_case a] g);; + +(* This is like the previous one except that it also rewrite on all + hypotheses except the ones given in the first argument. All the + modified hypotheses are generalized in the process and should be + introduced back later; the result is the pair of the tactic and the + list of hypotheses that have been generalized and cleared. *) +let mkDestructEq : + identifier list -> constr -> goal sigma -> tactic * identifier list = + fun not_on_hyp expr g -> + let hyps = pf_hyps g in + let to_revert = + Util.map_succeed + (fun (id,_,t) -> + if List.mem id not_on_hyp || not (Termops.occur_term expr t) + then failwith "is_expr_context"; + id) hyps in + let to_revert_constr = List.rev_map mkVar to_revert in + let type_of_expr = pf_type_of g expr in + let new_hyps = mkApp(delayed_force refl_equal, [|type_of_expr; expr|]):: + to_revert_constr in + tclTHENLIST + [h_generalize new_hyps; + (fun g2 -> + change_in_concl None + (pattern_occs [((false,[1]), expr)] (pf_env g2) Evd.empty (pf_concl g2)) g2); + simplest_case expr], to_revert + +let rec mk_intros_and_continue thin_intros (extra_eqn:bool) + cont_function (eqs:constr list) nb_lam (expr:constr) g = + observe_tac "mk_intros_and_continue" ( + let finalize () = if extra_eqn then + let teq = pf_get_new_id teq_id g in + tclTHENLIST + [ h_intro teq; + thin thin_intros; + h_intros thin_intros; + + tclMAP + (fun eq -> tclTRY (Equality.general_rewrite_in true all_occurrences (* deps proofs also: *) true teq eq false)) + (List.rev eqs); + (fun g1 -> + let ty_teq = pf_type_of g1 (mkVar teq) in + let teq_lhs,teq_rhs = + let _,args = try destApp ty_teq with _ -> Pp.msgnl (Printer.pr_goal (sig_it g1) ++ fnl () ++ pr_id teq ++ str ":" ++ Printer.pr_lconstr ty_teq); assert false in + args.(1),args.(2) + in + cont_function (mkVar teq::eqs) (replace_term teq_lhs teq_rhs expr) g1 + ) + ] + + else + tclTHENSEQ[ + thin thin_intros; + h_intros thin_intros; + cont_function eqs expr + ] + in + if nb_lam = 0 + then finalize () + else + match kind_of_term expr with + | Lambda (n, _, b) -> + let n1 = + match n with + Name x -> x + | Anonymous -> ano_id + in + let new_n = pf_get_new_id n1 g in + tclTHEN (h_intro new_n) + (mk_intros_and_continue thin_intros extra_eqn cont_function eqs + (pred nb_lam) (subst1 (mkVar new_n) b)) + | _ -> + assert false) g +(* finalize () *) +let const_of_ref = function + ConstRef kn -> kn + | _ -> anomaly "ConstRef expected" + +let simpl_iter clause = + reduce + (Lazy + {rBeta=true;rIota=true;rZeta= true; rDelta=false; + rConst = [ EvalConstRef (const_of_ref (delayed_force iter_ref))]}) +(* (Simpl (Some ([],mkConst (const_of_ref (delayed_force iter_ref))))) *) + clause + +(* The boolean value is_mes expresses that the termination is expressed + using a measure function instead of a well-founded relation. *) +let tclUSER tac is_mes l g = + let clear_tac = + match l with + | None -> h_clear true [] + | Some l -> tclMAP (fun id -> tclTRY (h_clear false [id])) (List.rev l) + in + tclTHENSEQ + [ + clear_tac; + if is_mes + then tclTHEN + (unfold_in_concl [(all_occurrences, evaluable_of_global_reference + (delayed_force ltof_ref))]) + tac + else tac + ] + g + + +let list_rewrite (rev:bool) (eqs: constr list) = + tclREPEAT + (List.fold_right + (fun eq i -> tclORELSE (rewriteLR eq) i) + (if rev then (List.rev eqs) else eqs) (tclFAIL 0 (mt())));; + +let base_leaf_terminate (func:global_reference) eqs expr = +(* let _ = msgnl (str "entering base_leaf") in *) + (fun g -> + let k',h = + match pf_get_new_ids [k_id;h_id] g with + [k';h] -> k',h + | _ -> assert false + in + tclTHENLIST + [observe_tac "first split" (split (ImplicitBindings [expr])); + observe_tac "second split" + (split (ImplicitBindings [delayed_force coq_O])); + observe_tac "intro k" (h_intro k'); + observe_tac "case on k" + (tclTHENS (simplest_case (mkVar k')) + [(tclTHEN (h_intro h) + (tclTHEN (simplest_elim (mkApp (delayed_force gt_antirefl, + [| delayed_force coq_O |]))) + default_auto)); tclIDTAC ]); + intros; + simpl_iter onConcl; + unfold_constr func; + list_rewrite true eqs; + default_auto] g);; + +(* La fonction est donnee en premier argument a la + fonctionnelle suivie d'autres Lambdas et de Case ... + Pour recuperer la fonction f a partir de la + fonctionnelle *) + +let get_f foncl = + match (kind_of_term (def_of_const foncl)) with + Lambda (Name f, _, _) -> f + |_ -> error "la fonctionnelle est mal definie";; + + +let rec compute_le_proofs = function + [] -> assumption + | a::tl -> + tclORELSE assumption + (tclTHENS + (fun g -> + let le_trans = delayed_force le_trans in + let t_le_trans = compute_renamed_type g le_trans in + let m_id = + let _,_,t = destProd t_le_trans in + let na,_,_ = destProd t in + Nameops.out_name na + in + apply_with_bindings + (le_trans, + ExplicitBindings[dummy_loc,NamedHyp m_id,a]) + g) + [compute_le_proofs tl; + tclORELSE (apply (delayed_force le_n)) assumption]) + +let make_lt_proof pmax le_proof = + tclTHENS + (fun g -> + let le_lt_trans = delayed_force le_lt_trans in + let t_le_lt_trans = compute_renamed_type g le_lt_trans in + let m_id = + let _,_,t = destProd t_le_lt_trans in + let na,_,_ = destProd t in + Nameops.out_name na + in + apply_with_bindings + (le_lt_trans, + ExplicitBindings[dummy_loc,NamedHyp m_id, pmax]) g) + [observe_tac "compute_le_proofs" (compute_le_proofs le_proof); + tclTHENLIST[observe_tac "lt_S_n" (apply (delayed_force lt_S_n)); default_full_auto]];; + +let rec list_cond_rewrite k def pmax cond_eqs le_proofs = + match cond_eqs with + [] -> tclIDTAC + | eq::eqs -> + (fun g -> + let t_eq = compute_renamed_type g (mkVar eq) in + let k_id,def_id = + let k_na,_,t = destProd t_eq in + let _,_,t = destProd t in + let def_na,_,_ = destProd t in + Nameops.out_name k_na,Nameops.out_name def_na + in + tclTHENS + (general_rewrite_bindings false all_occurrences + (* dep proofs also: *) true + (mkVar eq, + ExplicitBindings[dummy_loc, NamedHyp k_id, mkVar k; + dummy_loc, NamedHyp def_id, mkVar def]) false) + [list_cond_rewrite k def pmax eqs le_proofs; + observe_tac "make_lt_proof" (make_lt_proof pmax le_proofs)] g + ) + +let rec introduce_all_equalities func eqs values specs bound le_proofs + cond_eqs = + match specs with + [] -> + fun g -> + let ids = pf_ids_of_hyps g in + let s_max = mkApp(delayed_force coq_S, [|bound|]) in + let k = next_ident_away_in_goal k_id ids in + let ids = k::ids in + let h' = next_ident_away_in_goal (h'_id) ids in + let ids = h'::ids in + let def = next_ident_away_in_goal def_id ids in + tclTHENLIST + [observe_tac "introduce_all_equalities_final split" (split (ImplicitBindings [s_max])); + observe_tac "introduce_all_equalities_final intro k" (h_intro k); + tclTHENS + (observe_tac "introduce_all_equalities_final case k" (simplest_case (mkVar k))) + [ + tclTHENLIST[h_intro h'; + simplest_elim(mkApp(delayed_force lt_n_O,[|s_max|])); + default_full_auto]; + tclIDTAC + ]; + observe_tac "clearing k " (clear [k]); + observe_tac "intros k h' def" (h_intros [k;h';def]); + observe_tac "simple_iter" (simpl_iter onConcl); + observe_tac "unfold functional" + (unfold_in_concl[((true,[1]),evaluable_of_global_reference func)]); + observe_tac "rewriting equations" + (list_rewrite true eqs); + observe_tac ("cond rewrite "^(string_of_id k)) (list_cond_rewrite k def bound cond_eqs le_proofs); + observe_tac "refl equal" (apply (delayed_force refl_equal))] g + | spec1::specs -> + fun g -> + let ids = ids_of_named_context (pf_hyps g) in + let p = next_ident_away_in_goal p_id ids in + let ids = p::ids in + let pmax = next_ident_away_in_goal pmax_id ids in + let ids = pmax::ids in + let hle1 = next_ident_away_in_goal hle_id ids in + let ids = hle1::ids in + let hle2 = next_ident_away_in_goal hle_id ids in + let ids = hle2::ids in + let heq = next_ident_away_in_goal heq_id ids in + tclTHENLIST + [simplest_elim (mkVar spec1); + list_rewrite true eqs; + h_intros [p; heq]; + simplest_elim (mkApp(delayed_force max_constr, [| bound; mkVar p|])); + h_intros [pmax; hle1; hle2]; + introduce_all_equalities func eqs values specs + (mkVar pmax) ((mkVar pmax)::le_proofs) + (heq::cond_eqs)] g;; + +let string_match s = + if String.length s < 3 then failwith "string_match"; + try + for i = 0 to 3 do + if String.get s i <> String.get "Acc_" i then failwith "string_match" + done; + with Invalid_argument _ -> failwith "string_match" + +let retrieve_acc_var g = + (* Julien: I don't like this version .... *) + let hyps = pf_ids_of_hyps g in + map_succeed + (fun id -> string_match (string_of_id id);id) + hyps + +let rec introduce_all_values concl_tac is_mes acc_inv func context_fn + eqs hrec args values specs = + (match args with + [] -> + tclTHENLIST + [observe_tac "split" (split(ImplicitBindings + [context_fn (List.map mkVar (List.rev values))])); + observe_tac "introduce_all_equalities" (introduce_all_equalities func eqs + (List.rev values) (List.rev specs) (delayed_force coq_O) [] [])] + | arg::args -> + (fun g -> + let ids = ids_of_named_context (pf_hyps g) in + let rec_res = next_ident_away_in_goal rec_res_id ids in + let ids = rec_res::ids in + let hspec = next_ident_away_in_goal hspec_id ids in + let tac = + observe_tac "introduce_all_values" ( + introduce_all_values concl_tac is_mes acc_inv func context_fn eqs + hrec args + (rec_res::values)(hspec::specs)) in + (tclTHENS + (observe_tac "elim h_rec" + (simplest_elim (mkApp(mkVar hrec, Array.of_list arg))) + ) + [tclTHENLIST [h_intros [rec_res; hspec]; + tac]; + (tclTHENS + (observe_tac "acc_inv" (apply (Lazy.force acc_inv))) + [(* tclTHEN (tclTRY(list_rewrite true eqs)) *) + (observe_tac "h_assumption" h_assumption) + ; + tclTHENLIST + [ + tclTRY(list_rewrite true eqs); + observe_tac "user proof" + (fun g -> + tclUSER + concl_tac + is_mes + (Some (hrec::hspec::(retrieve_acc_var g)@specs)) + g + ) + ] + ] + ) + ]) g) + + ) + + +let rec_leaf_terminate f_constr concl_tac is_mes acc_inv hrec (func:global_reference) eqs expr = + match find_call_occs 0 f_constr expr with + | context_fn, args -> + observe_tac "introduce_all_values" + (introduce_all_values concl_tac is_mes acc_inv func context_fn eqs hrec args [] []) + +let proveterminate rec_arg_id is_mes acc_inv (hrec:identifier) + (f_constr:constr) (func:global_reference) base_leaf rec_leaf = + let rec proveterminate (eqs:constr list) (expr:constr) = + try + (* let _ = msgnl (str "entering proveterminate") in *) + let v = + match (kind_of_term expr) with + Case (ci, t, a, l) -> + (match find_call_occs 0 f_constr a with + _,[] -> + (fun g -> + let destruct_tac, rev_to_thin_intro = + mkDestructEq rec_arg_id a g in + tclTHENS destruct_tac + (list_map_i + (fun i -> mk_intros_and_continue + (List.rev rev_to_thin_intro) + true + proveterminate + eqs + ci.ci_cstr_nargs.(i)) + 0 (Array.to_list l)) g) + | _, _::_ -> + (match find_call_occs 0 f_constr expr with + _,[] -> observe_tac "base_leaf" (base_leaf func eqs expr) + | _, _:: _ -> + observe_tac "rec_leaf" + (rec_leaf is_mes acc_inv hrec func eqs expr))) + | _ -> + (match find_call_occs 0 f_constr expr with + _,[] -> + (try observe_tac "base_leaf" (base_leaf func eqs expr) + with e -> (msgerrnl (str "failure in base case");raise e )) + | _, _::_ -> + observe_tac "rec_leaf" + (rec_leaf is_mes acc_inv hrec func eqs expr)) in + v + with e -> begin msgerrnl(str "failure in proveterminate"); raise e end + in + proveterminate + +let hyp_terminates nb_args func = + let a_arrow_b = arg_type (constr_of_global func) in + let rev_args,b = decompose_prod_n nb_args a_arrow_b in + let left = + mkApp(delayed_force iter, + Array.of_list + (lift 5 a_arrow_b:: mkRel 3:: + constr_of_global func::mkRel 1:: + List.rev (list_map_i (fun i _ -> mkRel (6+i)) 0 rev_args) + ) + ) + in + let right = mkRel 5 in + let equality = mkApp(delayed_force eq, [|lift 5 b; left; right|]) in + let result = (mkProd ((Name def_id) , lift 4 a_arrow_b, equality)) in + let cond = mkApp(delayed_force lt, [|(mkRel 2); (mkRel 1)|]) in + let nb_iter = + mkApp(delayed_force ex, + [|delayed_force nat; + (mkLambda + (Name + p_id, + delayed_force nat, + (mkProd (Name k_id, delayed_force nat, + mkArrow cond result))))|])in + let value = mkApp(delayed_force coq_sig, + [|b; + (mkLambda (Name v_id, b, nb_iter))|]) in + compose_prod rev_args value + + + +let tclUSER_if_not_mes concl_tac is_mes names_to_suppress = + if is_mes + then tclCOMPLETE (h_simplest_apply (delayed_force well_founded_ltof)) + else tclUSER concl_tac is_mes names_to_suppress + +let termination_proof_header is_mes input_type ids args_id relation + rec_arg_num rec_arg_id tac wf_tac : tactic = + begin + fun g -> + let nargs = List.length args_id in + let pre_rec_args = + List.rev_map + mkVar (fst (list_chop (rec_arg_num - 1) args_id)) + in + let relation = substl pre_rec_args relation in + let input_type = substl pre_rec_args input_type in + let wf_thm = next_ident_away_in_goal (id_of_string ("wf_R")) ids in + let wf_rec_arg = + next_ident_away_in_goal + (id_of_string ("Acc_"^(string_of_id rec_arg_id))) + (wf_thm::ids) + in + let hrec = next_ident_away_in_goal hrec_id + (wf_rec_arg::wf_thm::ids) in + let acc_inv = + lazy ( + mkApp ( + delayed_force acc_inv_id, + [|input_type;relation;mkVar rec_arg_id|] + ) + ) + in + tclTHEN + (h_intros args_id) + (tclTHENS + (observe_tac + "first assert" + (assert_tac + (Name wf_rec_arg) + (mkApp (delayed_force acc_rel, + [|input_type;relation;mkVar rec_arg_id|]) + ) + ) + ) + [ + (* accesibility proof *) + tclTHENS + (observe_tac + "second assert" + (assert_tac + (Name wf_thm) + (mkApp (delayed_force well_founded,[|input_type;relation|])) + ) + ) + [ + (* interactive proof that the relation is well_founded *) + observe_tac "wf_tac" (wf_tac is_mes (Some args_id)); + (* this gives the accessibility argument *) + observe_tac + "apply wf_thm" + (h_simplest_apply (mkApp(mkVar wf_thm,[|mkVar rec_arg_id|])) + ) + ] + ; + (* rest of the proof *) + tclTHENSEQ + [observe_tac "generalize" + (onNLastHypsId (nargs+1) + (tclMAP (fun id -> + tclTHEN (h_generalize [mkVar id]) (h_clear false [id])) + )) + ; + observe_tac "h_fix" (h_fix (Some hrec) (nargs+1)); + h_intros args_id; + h_intro wf_rec_arg; + observe_tac "tac" (tac wf_rec_arg hrec acc_inv) + ] + ] + ) g + end + + + +let rec instantiate_lambda t l = + match l with + | [] -> t + | a::l -> + let (bound_name, _, body) = destLambda t in + instantiate_lambda (subst1 a body) l +;; + + +let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_arg_num : tactic = + begin + fun g -> + let ids = ids_of_named_context (pf_hyps g) in + let func_body = (def_of_const (constr_of_global func)) in + let (f_name, _, body1) = destLambda func_body in + let f_id = + match f_name with + | Name f_id -> next_ident_away_in_goal f_id ids + | Anonymous -> anomaly "Anonymous function" + in + let n_names_types,_ = decompose_lam_n nb_args body1 in + let n_ids,ids = + List.fold_left + (fun (n_ids,ids) (n_name,_) -> + match n_name with + | Name id -> + let n_id = next_ident_away_in_goal id ids in + n_id::n_ids,n_id::ids + | _ -> anomaly "anonymous argument" + ) + ([],(f_id::ids)) + n_names_types + in + let rec_arg_id = List.nth n_ids (rec_arg_num - 1) in + let expr = instantiate_lambda func_body (mkVar f_id::(List.map mkVar n_ids)) in + termination_proof_header + is_mes + input_type + ids + n_ids + relation + rec_arg_num + rec_arg_id + (fun rec_arg_id hrec acc_inv g -> + (proveterminate + [rec_arg_id] + is_mes + acc_inv + hrec + (mkVar f_id) + func + base_leaf_terminate + (rec_leaf_terminate (mkVar f_id) concl_tac) + [] + expr + ) + g + ) + (tclUSER_if_not_mes concl_tac) + g + end + +let get_current_subgoals_types () = + let pts = get_pftreestate () in + let _,subs = extract_open_pftreestate pts in + List.map snd ((* List.sort (fun (x,_) (y,_) -> x -y ) *)subs ) + +let build_and_l l = + let and_constr = Coqlib.build_coq_and () in + let conj_constr = coq_conj () in + let mk_and p1 p2 = + Term.mkApp(and_constr,[|p1;p2|]) in + let rec f = function + | [] -> failwith "empty list of subgoals!" + | [p] -> p,tclIDTAC,1 + | p1::pl -> + let c,tac,nb = f pl in + mk_and p1 c, + tclTHENS + (apply (constr_of_global conj_constr)) + [tclIDTAC; + tac + ],nb+1 + in f l + + +let is_rec_res id = + let rec_res_name = string_of_id rec_res_id in + let id_name = string_of_id id in + try + String.sub id_name 0 (String.length rec_res_name) = rec_res_name + with _ -> false + +let clear_goals = + let rec clear_goal t = + match kind_of_term t with + | Prod(Name id as na,t',b) -> + let b' = clear_goal b in + if noccurn 1 b' && (is_rec_res id) + then pop b' + else if b' == b then t + else mkProd(na,t',b') + | _ -> map_constr clear_goal t + in + List.map clear_goal + + +let build_new_goal_type () = + let sub_gls_types = get_current_subgoals_types () in + (* Pp.msgnl (str "sub_gls_types1 := " ++ Util.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *) + let sub_gls_types = clear_goals sub_gls_types in + (* Pp.msgnl (str "sub_gls_types2 := " ++ Util.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *) + let res = build_and_l sub_gls_types in + res + +let open_new_goal (build_proof:tactic -> tactic -> unit) using_lemmas ref_ goal_name (gls_type,decompose_and_tac,nb_goal) = + (* Pp.msgnl (str "gls_type := " ++ Printer.pr_lconstr gls_type); *) + let current_proof_name = get_current_proof_name () in + let name = match goal_name with + | Some s -> s + | None -> + try (add_suffix current_proof_name "_subproof") + with _ -> anomaly "open_new_goal with an unamed theorem" + in + let sign = Global.named_context () in + let sign = clear_proofs sign in + let na = next_global_ident_away name [] in + if occur_existential gls_type then + Util.error "\"abstract\" cannot handle existentials"; + let hook _ _ = + let opacity = + let na_ref = Libnames.Ident (dummy_loc,na) in + let na_global = Nametab.global na_ref in + match na_global with + ConstRef c -> + let cb = Global.lookup_constant c in + if cb.Declarations.const_opaque then true + else begin match cb.const_body with None -> true | _ -> false end + | _ -> anomaly "equation_lemma: not a constant" + in + let lemma = mkConst (Lib.make_con na) in + ref_ := Some lemma ; + let lid = ref [] in + let h_num = ref (-1) in + Flags.silently Vernacentries.interp (Vernacexpr.VernacAbort None); + build_proof + ( fun gls -> + let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in + tclTHENSEQ + [ + h_generalize [lemma]; + h_intro hid; + (fun g -> + let ids = pf_ids_of_hyps g in + tclTHEN + (Elim.h_decompose_and (mkVar hid)) + (fun g -> + let ids' = pf_ids_of_hyps g in + lid := List.rev (list_subtract ids' ids); + if !lid = [] then lid := [hid]; + tclIDTAC g + ) + g + ); + ] gls) + (fun g -> + match kind_of_term (pf_concl g) with + | App(f,_) when eq_constr f (well_founded ()) -> + Auto.h_auto None [] (Some []) g + | _ -> + incr h_num; + (observe_tac "finishing using" + ( + tclCOMPLETE( + tclFIRST[ + tclTHEN + (eapply_with_bindings (mkVar (List.nth !lid !h_num), NoBindings)) + e_assumption; + Eauto.eauto_with_bases + false + (true,5) + [delayed_force refl_equal] + [Auto.Hint_db.empty empty_transparent_state false] + ] + ) + ) + ) + g) +; + Lemmas.save_named opacity; + in + start_proof + na + (Decl_kinds.Global, Decl_kinds.Proof Decl_kinds.Lemma) + sign + gls_type + hook ; + if Indfun_common.is_strict_tcc () + then + by (tclIDTAC) + else + begin + by ( + fun g -> + tclTHEN + (decompose_and_tac) + (tclORELSE + (tclFIRST + (List.map + (fun c -> + tclTHENSEQ + [intros; + h_simplest_apply (interp_constr Evd.empty (Global.env()) c); + tclCOMPLETE Auto.default_auto + ] + ) + using_lemmas) + ) tclIDTAC) + g) + end; + try + by tclIDTAC; (* raises UserError _ if the proof is complete *) + if Flags.is_verbose () then (pp (Printer.pr_open_subgoals())) + with UserError _ -> + defined () + +;; + + +let com_terminate + tcc_lemma_name + tcc_lemma_ref + is_mes + fonctional_ref + input_type + relation + rec_arg_num + thm_name using_lemmas + nb_args + hook = + let start_proof (tac_start:tactic) (tac_end:tactic) = + let (evmap, env) = Lemmas.get_current_context() in + start_proof thm_name + (Global, Proof Lemma) (Environ.named_context_val env) + (hyp_terminates nb_args fonctional_ref) hook; + + by (observe_tac "starting_tac" tac_start); + by (observe_tac "whole_start" (whole_start tac_end nb_args is_mes fonctional_ref + input_type relation rec_arg_num )) + in + start_proof tclIDTAC tclIDTAC; + try + let new_goal_type = build_new_goal_type () in + open_new_goal start_proof using_lemmas tcc_lemma_ref + (Some tcc_lemma_name) + (new_goal_type); + + with Failure "empty list of subgoals!" -> + (* a non recursive function declared with measure ! *) + defined () + + + + +let ind_of_ref = function + | IndRef (ind,i) -> (ind,i) + | _ -> anomaly "IndRef expected" + +let (value_f:constr list -> global_reference -> constr) = + fun al fterm -> + let d0 = dummy_loc in + let rev_x_id_l = + ( + List.fold_left + (fun x_id_l _ -> + let x_id = next_ident_away_in_goal x_id x_id_l in + x_id::x_id_l + ) + [] + al + ) + in + let fun_body = + RCases + (d0,RegularStyle,None, + [RApp(d0, RRef(d0,fterm), List.rev_map (fun x_id -> RVar(d0, x_id)) rev_x_id_l), + (Anonymous,None)], + [d0, [v_id], [PatCstr(d0,(ind_of_ref + (delayed_force coq_sig_ref),1), + [PatVar(d0, Name v_id); + PatVar(d0, Anonymous)], + Anonymous)], + RVar(d0,v_id)]) + in + let value = + List.fold_left2 + (fun acc x_id a -> + RLambda + (d0, Name x_id, Explicit, RDynamic(d0, constr_in a), + acc + ) + ) + fun_body + rev_x_id_l + (List.rev al) + in + understand Evd.empty (Global.env()) value;; + +let (declare_fun : identifier -> logical_kind -> constr -> global_reference) = + fun f_id kind value -> + let ce = {const_entry_body = value; + const_entry_type = None; + const_entry_opaque = false; + const_entry_boxed = true} in + ConstRef(declare_constant f_id (DefinitionEntry ce, kind));; + +let (declare_f : identifier -> logical_kind -> constr list -> global_reference -> global_reference) = + fun f_id kind input_type fterm_ref -> + declare_fun f_id kind (value_f input_type fterm_ref);; + +let rec n_x_id ids n = + if n = 0 then [] + else let x = next_ident_away_in_goal x_id ids in + x::n_x_id (x::ids) (n-1);; + +let start_equation (f:global_reference) (term_f:global_reference) + (cont_tactic:identifier list -> tactic) g = + let ids = pf_ids_of_hyps g in + let terminate_constr = constr_of_global term_f in + let nargs = nb_prod (type_of_const terminate_constr) in + let x = n_x_id ids nargs in + tclTHENLIST [ + h_intros x; + unfold_in_concl [(all_occurrences, evaluable_of_global_reference f)]; + observe_tac "simplest_case" + (simplest_case (mkApp (terminate_constr, + Array.of_list (List.map mkVar x)))); + observe_tac "prove_eq" (cont_tactic x)] g;; + +let base_leaf_eq func eqs f_id g = + let ids = pf_ids_of_hyps g in + let k = next_ident_away_in_goal k_id ids in + let p = next_ident_away_in_goal p_id (k::ids) in + let v = next_ident_away_in_goal v_id (p::k::ids) in + let heq = next_ident_away_in_goal heq_id (v::p::k::ids) in + let heq1 = next_ident_away_in_goal heq_id (heq::v::p::k::ids) in + let hex = next_ident_away_in_goal hex_id (heq1::heq::v::p::k::ids) in + tclTHENLIST [ + h_intros [v; hex]; + simplest_elim (mkVar hex); + h_intros [p;heq1]; + tclTRY + (rewriteRL + (mkApp(mkVar heq1, + [|mkApp (delayed_force coq_S, [|mkVar p|]); + mkApp(delayed_force lt_n_Sn, [|mkVar p|]); f_id|]))); + simpl_iter onConcl; + tclTRY (unfold_in_concl [((true,[1]), evaluable_of_global_reference func)]); + observe_tac "list_revrite" (list_rewrite true eqs); + apply (delayed_force refl_equal)] g;; + +let f_S t = mkApp(delayed_force coq_S, [|t|]);; + + +let rec introduce_all_values_eq cont_tac functional termine + f p heq1 pmax bounds le_proofs eqs ids = + function + [] -> + let heq2 = next_ident_away_in_goal heq_id ids in + tclTHENLIST + [pose_proof (Name heq2) + (mkApp(mkVar heq1, [|f_S(f_S(mkVar pmax))|])); + simpl_iter (onHyp heq2); + unfold_in_hyp [((true,[1]), evaluable_of_global_reference + (global_of_constr functional))] + (heq2, InHyp); + tclTHENS + (fun gls -> + let t_eq = compute_renamed_type gls (mkVar heq2) in + let def_id = + let _,_,t = destProd t_eq in let def_na,_,_ = destProd t in + Nameops.out_name def_na + in + observe_tac "rewrite heq" (general_rewrite_bindings false all_occurrences + (* dep proofs also: *) true (mkVar heq2, + ExplicitBindings[dummy_loc,NamedHyp def_id, + f]) false) gls) + [tclTHENLIST + [observe_tac "list_rewrite" (list_rewrite true eqs); + cont_tac pmax le_proofs]; + tclTHENLIST[apply (delayed_force le_lt_SS); + compute_le_proofs le_proofs]]] + | arg::args -> + let v' = next_ident_away_in_goal v_id ids in + let ids = v'::ids in + let hex' = next_ident_away_in_goal hex_id ids in + let ids = hex'::ids in + let p' = next_ident_away_in_goal p_id ids in + let ids = p'::ids in + let new_pmax = next_ident_away_in_goal pmax_id ids in + let ids = pmax::ids in + let hle1 = next_ident_away_in_goal hle_id ids in + let ids = hle1::ids in + let hle2 = next_ident_away_in_goal hle_id ids in + let ids = hle2::ids in + let heq = next_ident_away_in_goal heq_id ids in + let ids = heq::ids in + let heq2 = next_ident_away_in_goal heq_id ids in + let ids = heq2::ids in + tclTHENLIST + [mkCaseEq(mkApp(termine, Array.of_list arg)); + h_intros [v'; hex']; + simplest_elim(mkVar hex'); + h_intros [p']; + simplest_elim(mkApp(delayed_force max_constr, [|mkVar pmax; + mkVar p'|])); + h_intros [new_pmax;hle1;hle2]; + introduce_all_values_eq + (fun pmax' le_proofs'-> + tclTHENLIST + [cont_tac pmax' le_proofs'; + h_intros [heq;heq2]; + observe_tac ("rewriteRL " ^ (string_of_id heq2)) + (tclTRY (rewriteLR (mkVar heq2))); + tclTRY (tclTHENS + ( fun g -> + let t_eq = compute_renamed_type g (mkVar heq) in + let k_id,def_id = + let k_na,_,t = destProd t_eq in + let _,_,t = destProd t in + let def_na,_,_ = destProd t in + Nameops.out_name k_na,Nameops.out_name def_na + in + let c_b = (mkVar heq, + ExplicitBindings + [dummy_loc, NamedHyp k_id, + f_S(mkVar pmax'); + dummy_loc, NamedHyp def_id, f]) + in + observe_tac "general_rewrite_bindings" ( (general_rewrite_bindings false all_occurrences (* dep proofs also: *) true + c_b false)) + g + ) + [tclIDTAC; + tclTHENLIST + [apply (delayed_force le_lt_n_Sm); + compute_le_proofs le_proofs']])]) + functional termine f p heq1 new_pmax + (p'::bounds)((mkVar pmax)::le_proofs) eqs + (heq2::heq::hle2::hle1::new_pmax::p'::hex'::v'::ids) args] + + +let rec_leaf_eq termine f ids functional eqs expr fn args = + let p = next_ident_away_in_goal p_id ids in + let ids = p::ids in + let v = next_ident_away_in_goal v_id ids in + let ids = v::ids in + let hex = next_ident_away_in_goal hex_id ids in + let ids = hex::ids in + let heq1 = next_ident_away_in_goal heq_id ids in + let ids = heq1::ids in + let hle1 = next_ident_away_in_goal hle_id ids in + let ids = hle1::ids in + tclTHENLIST + [observe_tac "intros v hex" (h_intros [v;hex]); + simplest_elim (mkVar hex); + h_intros [p;heq1]; + h_generalize [mkApp(delayed_force le_n,[|mkVar p|])]; + h_intros [hle1]; + observe_tac "introduce_all_values_eq" (introduce_all_values_eq + (fun _ _ -> tclIDTAC) + functional termine f p heq1 p [] [] eqs ids args); + observe_tac "failing here" (apply (delayed_force refl_equal))] + +let rec prove_eq (termine:constr) (f:constr)(functional:global_reference) + (eqs:constr list) (expr:constr) = +(* tclTRY *) + observe_tac "prove_eq" (match kind_of_term expr with + Case(ci,t,a,l) -> + (match find_call_occs 0 f a with + _,[] -> + (fun g -> + let destruct_tac,rev_to_thin_intro = mkDestructEq [] a g in + tclTHENS + destruct_tac + (list_map_i + (fun i -> mk_intros_and_continue + (List.rev rev_to_thin_intro) true + (prove_eq termine f functional) + eqs ci.ci_cstr_nargs.(i)) + 0 (Array.to_list l)) g) + | _,_::_ -> + (match find_call_occs 0 f expr with + _,[] -> observe_tac "base_leaf_eq(1)" (base_leaf_eq functional eqs f) + | fn,args -> + fun g -> + let ids = ids_of_named_context (pf_hyps g) in + observe_tac "rec_leaf_eq" (rec_leaf_eq termine f ids + (constr_of_global functional) + eqs expr fn args) g)) + | _ -> + (match find_call_occs 0 f expr with + _,[] -> observe_tac "base_leaf_eq(2)" ( base_leaf_eq functional eqs f) + | fn,args -> + fun g -> + let ids = ids_of_named_context (pf_hyps g) in + observe_tac "rec_leaf_eq" (rec_leaf_eq + termine f ids (constr_of_global functional) + eqs expr fn args) g));; + +let (com_eqn : identifier -> + global_reference -> global_reference -> global_reference + -> constr -> unit) = + fun eq_name functional_ref f_ref terminate_ref equation_lemma_type -> + let opacity = + match terminate_ref with + | ConstRef c -> + let cb = Global.lookup_constant c in + if cb.Declarations.const_opaque then true + else begin match cb.const_body with None -> true | _ -> false end + | _ -> anomaly "terminate_lemma: not a constant" + in + let (evmap, env) = Lemmas.get_current_context() in + let f_constr = (constr_of_global f_ref) in + let equation_lemma_type = subst1 f_constr equation_lemma_type in + (start_proof eq_name (Global, Proof Lemma) + (Environ.named_context_val env) equation_lemma_type (fun _ _ -> ()); + by + (start_equation f_ref terminate_ref + (fun x -> + prove_eq + (constr_of_global terminate_ref) + f_constr + functional_ref + [] + (instantiate_lambda + (def_of_const (constr_of_global functional_ref)) + (f_constr::List.map mkVar x) + ) + ) + ); +(* (try Vernacentries.interp (Vernacexpr.VernacShow Vernacexpr.ShowProof) with _ -> ()); *) +(* Vernacentries.interp (Vernacexpr.VernacShow Vernacexpr.ShowScript); *) + Flags.silently (fun () -> Lemmas.save_named opacity) () ; +(* Pp.msgnl (str "eqn finished"); *) + + );; + +let nf_zeta env = + Reductionops.clos_norm_flags (Closure.RedFlags.mkflags [Closure.RedFlags.fZETA]) + env + Evd.empty + +let nf_betaiotazeta = (* Reductionops.local_strong Reductionops.whd_betaiotazeta *) + let clos_norm_flags flgs env sigma t = + Closure.norm_val (Closure.create_clos_infos flgs env) (Closure.inject (Reductionops.nf_evar sigma t)) in + clos_norm_flags Closure.betaiotazeta Environ.empty_env Evd.empty + + +let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num eq + generate_induction_principle using_lemmas : unit = + let function_type = interp_constr Evd.empty (Global.env()) type_of_f in + let env = push_named (function_name,None,function_type) (Global.env()) in +(* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *) + let equation_lemma_type = + nf_betaiotazeta + (interp_gen (OfType None) Evd.empty env ~impls:rec_impls eq) + in +(* Pp.msgnl (str "lemma type := " ++ Printer.pr_lconstr equation_lemma_type ++ fnl ()); *) + let res_vars,eq' = decompose_prod equation_lemma_type in + let env_eq' = Environ.push_rel_context (List.map (fun (x,y) -> (x,None,y)) res_vars) env in + let eq' = nf_zeta env_eq' eq' in + let res = +(* Pp.msgnl (str "res_var :=" ++ Printer.pr_lconstr_env (push_rel_context (List.map (function (x,t) -> (x,None,t)) res_vars) env) eq'); *) +(* Pp.msgnl (str "rec_arg_num := " ++ str (string_of_int rec_arg_num)); *) +(* Pp.msgnl (str "eq' := " ++ str (string_of_int rec_arg_num)); *) + match kind_of_term eq' with + | App(e,[|_;_;eq_fix|]) -> + mkLambda (Name function_name,function_type,subst_var function_name (compose_lam res_vars eq_fix)) + | _ -> failwith "Recursive Definition (res not eq)" + in + let pre_rec_args,function_type_before_rec_arg = decompose_prod_n (rec_arg_num - 1) function_type in + let (_, rec_arg_type, _) = destProd function_type_before_rec_arg in + let arg_types = List.rev_map snd (fst (decompose_prod_n (List.length res_vars) function_type)) in + let equation_id = add_suffix function_name "_equation" in + let functional_id = add_suffix function_name "_F" in + let term_id = add_suffix function_name "_terminate" in + let functional_ref = declare_fun functional_id (IsDefinition Definition) res in + let env_with_pre_rec_args = push_rel_context(List.map (function (x,t) -> (x,None,t)) pre_rec_args) env in + let relation = + interp_constr + Evd.empty + env_with_pre_rec_args + r + in + let tcc_lemma_name = add_suffix function_name "_tcc" in + let tcc_lemma_constr = ref None in + (* let _ = Pp.msgnl (str "relation := " ++ Printer.pr_lconstr_env env_with_pre_rec_args relation) in *) + let hook _ _ = + let term_ref = Nametab.locate (qualid_of_ident term_id) in + let f_ref = declare_f function_name (IsProof Lemma) arg_types term_ref in +(* message "start second proof"; *) + let stop = ref false in + begin + try com_eqn equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type) + with e -> + begin + if Tacinterp.get_debug () <> Tactic_debug.DebugOff + then pperrnl (str "Cannot create equation Lemma " ++ Cerrors.explain_exn e) + else anomaly "Cannot create equation Lemma" + ; +(* ignore(try Vernacentries.vernac_reset_name (Util.dummy_loc,functional_id) with _ -> ()); *) + stop := true; + end + end; + if not !stop + then + let eq_ref = Nametab.locate (qualid_of_ident equation_id ) in + let f_ref = destConst (constr_of_global f_ref) + and functional_ref = destConst (constr_of_global functional_ref) + and eq_ref = destConst (constr_of_global eq_ref) in + generate_induction_principle f_ref tcc_lemma_constr + functional_ref eq_ref rec_arg_num rec_arg_type (nb_prod res) relation; + if Flags.is_verbose () + then msgnl (h 1 (Ppconstr.pr_id function_name ++ + spc () ++ str"is defined" )++ fnl () ++ + h 1 (Ppconstr.pr_id equation_id ++ + spc () ++ str"is defined" ) + ) + in + try + com_terminate + tcc_lemma_name + tcc_lemma_constr + is_mes functional_ref + rec_arg_type + relation rec_arg_num + term_id + using_lemmas + (List.length res_vars) + hook + with e -> + begin + ignore(try Vernacentries.vernac_reset_name (Util.dummy_loc,functional_id) with _ -> ()); +(* anomaly "Cannot create termination Lemma" *) + raise e + end + + + diff --git a/plugins/funind/recdef_plugin.mllib b/plugins/funind/recdef_plugin.mllib new file mode 100644 index 00000000..31818c39 --- /dev/null +++ b/plugins/funind/recdef_plugin.mllib @@ -0,0 +1,11 @@ +Indfun_common +Rawtermops +Recdef +Rawterm_to_relation +Functional_principles_proofs +Functional_principles_types +Invfun +Indfun +Merge +G_indfun +Recdef_plugin_mod diff --git a/plugins/funind/vo.itarget b/plugins/funind/vo.itarget new file mode 100644 index 00000000..33c96830 --- /dev/null +++ b/plugins/funind/vo.itarget @@ -0,0 +1 @@ +Recdef.vo -- cgit v1.2.3