summaryrefslogtreecommitdiff
path: root/plugins/funind/invfun.ml
diff options
context:
space:
mode:
Diffstat (limited to 'plugins/funind/invfun.ml')
-rw-r--r--plugins/funind/invfun.ml323
1 files changed, 155 insertions, 168 deletions
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index 26fc88a6..bed95740 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -1,17 +1,21 @@
(************************************************************************)
-(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2016 *)
+(* * The Coq Proof Assistant / The Coq Development Team *)
+(* v * INRIA, CNRS and contributors - Copyright 1999-2018 *)
+(* <O___,, * (see CREDITS file for the list of authors) *)
(* \VV/ **************************************************************)
-(* // * This file is distributed under the terms of the *)
-(* * GNU Lesser General Public License Version 2.1 *)
+(* // * This file is distributed under the terms of the *)
+(* * GNU Lesser General Public License Version 2.1 *)
+(* * (see LICENSE file for the text of the license) *)
(************************************************************************)
-open Tacexpr
+open Ltac_plugin
open Declarations
open CErrors
open Util
open Names
open Term
+open Constr
+open EConstr
open Vars
open Pp
open Globnames
@@ -23,30 +27,7 @@ open Misctypes
open Termops
open Context.Rel.Declaration
-(* Some pretty printing function for debugging purpose *)
-
-let pr_binding prc =
- function
- | loc, NamedHyp id, c -> hov 1 (Ppconstr.pr_id id ++ str " := " ++ Pp.cut () ++ prc c)
- | loc, AnonHyp n, c -> hov 1 (int n ++ str " := " ++ Pp.cut () ++ prc c)
-
-let pr_bindings prc prlc = function
- | ImplicitBindings l ->
- brk (1,1) ++ str "with" ++ brk (1,1) ++
- pr_sequence prc l
- | ExplicitBindings l ->
- brk (1,1) ++ str "with" ++ brk (1,1) ++
- pr_sequence (fun b -> str"(" ++ pr_binding prlc b ++ str")") l
- | 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)
+module RelDecl = Context.Rel.Declaration
(* The local debugging mechanism *)
(* let msgnl = Pp.msgnl *)
@@ -77,12 +58,6 @@ let do_observe_tac s tac g =
CErrors.iprint e ++ str " on goal" ++ fnl() ++ goal ));
iraise reraise;;
-
-let observe_tac_strm s tac g =
- if do_observe ()
- then do_observe_tac s tac g
- else tac g
-
let observe_tac s tac g =
if do_observe ()
then do_observe_tac (str s) tac g
@@ -106,12 +81,8 @@ let thin ids gl = Proofview.V82.of_tactic (Tactics.clear ids) gl
let make_eq () =
try
- Universes.constr_of_global (Coqlib.build_coq_eq ())
+ EConstr.of_constr (Universes.constr_of_global (Coqlib.build_coq_eq ()))
with _ -> assert false
-let make_eq_refl () =
- try
- Universes.constr_of_global (Coqlib.build_coq_eq_refl ())
- with _ -> assert false
(* [generate_type g_to_f f graph i] build the completeness (resp. correctness) lemma type if [g_to_f = true]
@@ -129,15 +100,16 @@ let make_eq_refl () =
let generate_type evd 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 evd',graph =
- Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd graph)))
+ Evd.fresh_global (Global.env ()) !evd (Globnames.IndRef (fst (destInd !evd graph)))
in
+ let graph = EConstr.of_constr graph in
evd:=evd';
let graph_arity = Typing.e_type_of (Global.env ()) evd graph in
- let ctxt,_ = decompose_prod_assum graph_arity in
+ let ctxt,_ = decompose_prod_assum !evd graph_arity in
let fun_ctxt,res_type =
match ctxt with
- | [] | [_] -> anomaly (Pp.str "Not a valid context")
- | decl :: fun_ctxt -> fun_ctxt, get_type decl
+ | [] | [_] -> anomaly (Pp.str "Not a valid context.")
+ | decl :: fun_ctxt -> fun_ctxt, RelDecl.get_type decl
in
let rec args_from_decl i accu = function
| [] -> accu
@@ -148,13 +120,13 @@ let generate_type evd g_to_f f graph i =
args_from_decl (succ i) (t :: accu) l
in
(*i We need to name the vars [res] and [fv] i*)
- let filter = fun decl -> match get_name decl with
+ let filter = fun decl -> match RelDecl.get_name decl with
| Name id -> Some id
| Anonymous -> None
in
- let named_ctxt = List.map_filter filter fun_ctxt in
+ let named_ctxt = Id.Set.of_list (List.map_filter filter fun_ctxt) in
let res_id = Namegen.next_ident_away_in_goal (Id.of_string "_res") named_ctxt in
- let fv_id = Namegen.next_ident_away_in_goal (Id.of_string "fv") (res_id :: named_ctxt) in
+ let fv_id = Namegen.next_ident_away_in_goal (Id.of_string "fv") (Id.Set.add res_id named_ctxt) in
(*i we can then type the argument to be applied to the function [f] i*)
let args_as_rels = Array.of_list (args_from_decl 1 [] fun_ctxt) in
(*i
@@ -191,15 +163,16 @@ let generate_type evd g_to_f f graph i =
WARNING: while convertible, [type_of body] and [type] can be non equal
*)
let find_induction_principle evd f =
- let f_as_constant,u = match kind_of_term f with
+ let f_as_constant,u = match EConstr.kind !evd f with
| Const c' -> c'
- | _ -> error "Must be used with a function"
+ | _ -> user_err Pp.(str "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 evd',rect_lemma = Evd.fresh_global (Global.env ()) !evd (Globnames.ConstRef rect_lemma) in
+ let rect_lemma = EConstr.of_constr rect_lemma in
let evd',typ = Typing.type_of ~refresh:true (Global.env ()) evd' rect_lemma in
evd:=evd';
rect_lemma,typ
@@ -209,14 +182,13 @@ let rec generate_fresh_id x avoid i =
if i == 0
then []
else
- let id = Namegen.next_ident_away_in_goal x avoid in
+ let id = Namegen.next_ident_away_in_goal x (Id.Set.of_list 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 ]
+(* [prove_fun_correct 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.
@@ -237,29 +209,29 @@ let rec generate_fresh_id x avoid i =
\end{enumerate}
*)
-let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes lemmas_types_infos i : tactic =
+let prove_fun_correct evd funs_constr graphs_constr schemes lemmas_types_infos i : Tacmach.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\]
*)
(* we the get the definition of the graphs block *)
- let graph_ind,u = destInd graphs_constr.(i) in
+ let graph_ind,u = destInd evd 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
+ let princ_infos = Tactics.compute_elim_sig evd princ_type in
(* The number of args of the function is then easily computable *)
- let nb_fun_args = nb_prod (pf_concl g) - 2 in
+ let nb_fun_args = nb_prod (project g) (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 functional principle is defined in the
environment 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 principle_id = Namegen.next_ident_away_in_goal (Id.of_string "princ") (Id.Set.of_list ids) in
let ids = principle_id :: ids in
(* We get the branches of the principle *)
let branches = List.rev princ_infos.branches in
@@ -268,14 +240,14 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
List.map
(fun decl ->
List.map
- (fun id -> Loc.ghost, IntroNaming (IntroIdentifier id))
- (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum (get_type decl)))))
+ (fun id -> CAst.make @@ IntroNaming (IntroIdentifier id))
+ (generate_fresh_id (Id.of_string "y") ids (List.length (fst (decompose_prod_assum evd (RelDecl.get_type decl)))))
)
branches
in
(* before building the full intro pattern for the principle *)
let eq_ind = make_eq () in
- let eq_construct = mkConstructUi (destInd eq_ind, 1) in
+ let eq_construct = mkConstructUi (destInd evd 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
@@ -284,10 +256,10 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
(* We get the identifiers of this branch *)
let pre_args =
List.fold_right
- (fun (_,pat) acc ->
+ (fun {CAst.v=pat} acc ->
match pat with
| IntroNaming (IntroIdentifier id) -> id::acc
- | _ -> anomaly (Pp.str "Not an identifier")
+ | _ -> anomaly (Pp.str "Not an identifier.")
)
(List.nth intro_pats (pred i))
[]
@@ -304,17 +276,18 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
List.fold_right
(fun hid acc ->
let type_of_hid = pf_unsafe_type_of g (mkVar hid) in
- match kind_of_term type_of_hid with
+ let sigma = project g in
+ match EConstr.kind sigma type_of_hid with
| Prod(_,_,t') ->
begin
- match kind_of_term t' with
+ match EConstr.kind sigma t' with
| Prod(_,t'',t''') ->
begin
- match kind_of_term t'',kind_of_term t''' with
+ match EConstr.kind sigma t'',EConstr.kind sigma t''' with
| App(eq,args), App(graph',_)
when
- (eq_constr eq eq_ind) &&
- Array.exists (Constr.eq_constr_nounivs graph') graphs_constr ->
+ (EConstr.eq_constr sigma eq eq_ind) &&
+ Array.exists (EConstr.eq_constr_nounivs sigma graph') graphs_constr ->
(args.(2)::(mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);args.(2)|]))|]))
::acc)
| _ -> mkVar hid :: acc
@@ -360,7 +333,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
in
(* observe (str "constructor := " ++ Printer.pr_lconstr_env (pf_env g) app_constructor); *)
(
- tclTHENSEQ
+ tclTHENLIST
[
observe_tac("h_intro_patterns ") (let l = (List.nth intro_pats (pred i)) in
match l with
@@ -379,7 +352,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
Locusops.onConcl);
observe_tac ("toto ") tclIDTAC;
- (* introducing the the result of the graph and the equality hypothesis *)
+ (* introducing the result of the graph and the equality hypothesis *)
observe_tac "introducing" (tclMAP (fun x -> Proofview.V82.of_tactic (Simple.intro x)) [res;hres]);
(* replacing [res] with its value *)
observe_tac "rewriting res value" (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar hres)));
@@ -395,11 +368,11 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
Array.map
(fun ((_,(ctxt,concl))) ->
match ctxt with
- | [] | [_] | [_;_] -> anomaly (Pp.str "bad context")
+ | [] | [_] | [_;_] -> anomaly (Pp.str "bad context.")
| hres::res::decl::ctxt ->
- let res = Termops.it_mkLambda_or_LetIn
- (Termops.it_mkProd_or_LetIn concl [hres;res])
- (LocalAssum (get_name decl, get_type decl) :: ctxt)
+ let res = EConstr.it_mkLambda_or_LetIn
+ (EConstr.it_mkProd_or_LetIn concl [hres;res])
+ (LocalAssum (RelDecl.get_name decl, RelDecl.get_type decl) :: ctxt)
in
res
)
@@ -415,7 +388,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
let params_bindings,avoid =
List.fold_left2
(fun (bindings,avoid) decl p ->
- let id = Namegen.next_ident_away (Nameops.out_name (get_name decl)) avoid in
+ let id = Namegen.next_ident_away (Nameops.Name.get_id (RelDecl.get_name decl)) (Id.Set.of_list avoid) in
p::bindings,id::avoid
)
([],pf_ids_of_hyps g)
@@ -425,7 +398,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
let lemmas_bindings =
List.rev (fst (List.fold_left2
(fun (bindings,avoid) decl p ->
- let id = Namegen.next_ident_away (Nameops.out_name (get_name decl)) avoid in
+ let id = Namegen.next_ident_away (Nameops.Name.get_id (RelDecl.get_name decl)) (Id.Set.of_list avoid) in
(nf_zeta p)::bindings,id::avoid)
([],avoid)
princ_infos.predicates
@@ -433,7 +406,7 @@ let prove_fun_correct evd functional_induction funs_constr graphs_constr schemes
in
(params_bindings@lemmas_bindings)
in
- tclTHENSEQ
+ tclTHENLIST
[
observe_tac "principle" (Proofview.V82.of_tactic (assert_by
(Name principle_id)
@@ -465,7 +438,7 @@ let generalize_dependent_of x hyp g =
tclMAP
(function
| LocalAssum (id,t) when not (Id.equal id hyp) &&
- (Termops.occur_var (pf_env g) x t) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id])
+ (Termops.occur_var (pf_env g) (project g) x t) -> tclTHEN (Proofview.V82.of_tactic (Tactics.generalize [mkVar id])) (thin [id])
| _ -> tclIDTAC
)
(pf_hyps g)
@@ -486,46 +459,47 @@ let tauto =
let rec intros_with_rewrite g =
observe_tac "intros_with_rewrite" intros_with_rewrite_aux g
-and intros_with_rewrite_aux : tactic =
+and intros_with_rewrite_aux : Tacmach.tactic =
fun g ->
let eq_ind = make_eq () in
- match kind_of_term (pf_concl g) with
+ let sigma = project g in
+ match EConstr.kind sigma (pf_concl g) with
| Prod(_,t,t') ->
begin
- match kind_of_term t with
- | App(eq,args) when (eq_constr eq eq_ind) ->
+ match EConstr.kind sigma t with
+ | App(eq,args) when (EConstr.eq_constr sigma 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 [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g
- else if isVar args.(1) && (Environ.evaluable_named (destVar args.(1)) (pf_env g))
- then tclTHENSEQ[
- Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))]);
- tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))] ((destVar args.(1)),Locus.InHyp) )))
+ tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id); thin [id]; intros_with_rewrite ] g
+ else if isVar sigma args.(1) && (Environ.evaluable_named (destVar sigma args.(1)) (pf_env g))
+ then tclTHENLIST[
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))]);
+ tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(1)))] ((destVar sigma args.(1)),Locus.InHyp) )))
(pf_ids_of_hyps g);
intros_with_rewrite
] g
- else if isVar args.(2) && (Environ.evaluable_named (destVar args.(2)) (pf_env g))
- then tclTHENSEQ[
- Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))]);
- tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))] ((destVar args.(2)),Locus.InHyp) )))
+ else if isVar sigma args.(2) && (Environ.evaluable_named (destVar sigma args.(2)) (pf_env g))
+ then tclTHENLIST[
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))]);
+ tclMAP (fun id -> tclTRY(Proofview.V82.of_tactic (unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar sigma args.(2)))] ((destVar sigma args.(2)),Locus.InHyp) )))
(pf_ids_of_hyps g);
intros_with_rewrite
] g
- else if isVar args.(1)
+ else if isVar sigma args.(1)
then
let id = pf_get_new_id (Id.of_string "y") g in
- tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id);
- generalize_dependent_of (destVar args.(1)) id;
+ tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id);
+ generalize_dependent_of (destVar sigma args.(1)) id;
tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id)));
intros_with_rewrite
]
g
- else if isVar args.(2)
+ else if isVar sigma args.(2)
then
let id = pf_get_new_id (Id.of_string "y") g in
- tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id);
- generalize_dependent_of (destVar args.(2)) id;
+ tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id);
+ generalize_dependent_of (destVar sigma args.(2)) id;
tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar id)));
intros_with_rewrite
]
@@ -533,21 +507,21 @@ and intros_with_rewrite_aux : tactic =
else
begin
let id = pf_get_new_id (Id.of_string "y") g in
- tclTHENSEQ[
+ tclTHENLIST[
Proofview.V82.of_tactic (Simple.intro id);
tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id)));
intros_with_rewrite
] g
end
- | Ind _ when eq_constr t (Coqlib.build_coq_False ()) ->
+ | Ind _ when EConstr.eq_constr sigma t (EConstr.of_constr (Universes.constr_of_global @@ Coqlib.build_coq_False ())) ->
Proofview.V82.of_tactic tauto g
| Case(_,_,v,_) ->
- tclTHENSEQ[
+ tclTHENLIST[
Proofview.V82.of_tactic (simplest_case v);
intros_with_rewrite
] g
| LetIn _ ->
- tclTHENSEQ[
+ tclTHENLIST[
Proofview.V82.of_tactic (reduce
(Genredexpr.Cbv
{Redops.all_flags
@@ -559,10 +533,10 @@ and intros_with_rewrite_aux : tactic =
] g
| _ ->
let id = pf_get_new_id (Id.of_string "y") g in
- tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id);intros_with_rewrite] g
+ tclTHENLIST [ Proofview.V82.of_tactic (Simple.intro id);intros_with_rewrite] g
end
| LetIn _ ->
- tclTHENSEQ[
+ tclTHENLIST[
Proofview.V82.of_tactic (reduce
(Genredexpr.Cbv
{Redops.all_flags
@@ -577,9 +551,9 @@ and intros_with_rewrite_aux : tactic =
let rec reflexivity_with_destruct_cases g =
let destruct_case () =
try
- match kind_of_term (snd (destApp (pf_concl g))).(2) with
+ match EConstr.kind (project g) (snd (destApp (project g) (pf_concl g))).(2) with
| Case(_,_,v,_) ->
- tclTHENSEQ[
+ tclTHENLIST[
Proofview.V82.of_tactic (simplest_case v);
Proofview.V82.of_tactic intros;
observe_tac "reflexivity_with_destruct_cases" reflexivity_with_destruct_cases
@@ -588,18 +562,23 @@ let rec reflexivity_with_destruct_cases g =
with e when CErrors.noncritical e -> Proofview.V82.of_tactic reflexivity
in
let eq_ind = make_eq () in
+ let my_inj_flags = Some {
+ Equality.keep_proof_equalities = false;
+ injection_in_context = false; (* for compatibility, necessary *)
+ injection_pattern_l2r_order = false; (* probably does not matter; except maybe with dependent hyps *)
+ } in
let discr_inject =
Tacticals.onAllHypsAndConcl (
fun sc g ->
match sc with
None -> tclIDTAC g
| Some id ->
- match kind_of_term (pf_unsafe_type_of g (mkVar id)) with
- | App(eq,[|_;t1;t2|]) when eq_constr eq eq_ind ->
+ match EConstr.kind (project g) (pf_unsafe_type_of g (mkVar id)) with
+ | App(eq,[|_;t1;t2|]) when EConstr.eq_constr (project g) eq eq_ind ->
if Equality.discriminable (pf_env g) (project g) t1 t2
then Proofview.V82.of_tactic (Equality.discrHyp id) g
- else if Equality.injectable (pf_env g) (project g) t1 t2
- then tclTHENSEQ [Proofview.V82.of_tactic (Equality.injHyp None id);thin [id];intros_with_rewrite] g
+ else if Equality.injectable (pf_env g) (project g) ~keep_proofs:None t1 t2
+ then tclTHENLIST [Proofview.V82.of_tactic (Equality.injHyp my_inj_flags None id);thin [id];intros_with_rewrite] g
else tclIDTAC g
| _ -> tclIDTAC g
)
@@ -646,25 +625,25 @@ let rec reflexivity_with_destruct_cases g =
*)
-let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
+let prove_fun_complete funcs graphs schemes lemmas_types_infos i : Tacmach.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 concl ctxt))
+ (fun (_,(ctxt,concl)) -> nf_zeta (EConstr.it_mkLambda_or_LetIn 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 graph_principle = nf_zeta (EConstr.of_constr schemes.(i)) in
let princ_type = pf_unsafe_type_of g graph_principle in
- let princ_infos = Tactics.compute_elim_sig princ_type in
+ let princ_infos = Tactics.compute_elim_sig (project g) 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 nb_fun_args = nb_prod (project g) (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) *)
@@ -682,7 +661,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
(fun decl ->
List.map
(fun id -> id)
- (generate_fresh_id (Id.of_string "y") ids (nb_prod (get_type decl)))
+ (generate_fresh_id (Id.of_string "y") ids (nb_prod (project g) (RelDecl.get_type decl)))
)
branches
in
@@ -690,20 +669,20 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
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 rewrite_tac j ids : Tacmach.tactic =
let graph_def = graphs.(j) in
let infos =
- try find_Function_infos (fst (destConst funcs.(j)))
- with Not_found -> error "No graph found"
+ try find_Function_infos (fst (destConst (project g) funcs.(j)))
+ with Not_found -> user_err Pp.(str "No graph found")
in
if infos.is_general
|| Rtree.is_infinite Declareops.eq_recarg graph_def.mind_recargs
then
let eq_lemma =
try Option.get (infos).equation_lemma
- with Option.IsNone -> anomaly (Pp.str "Cannot find equation lemma")
+ with Option.IsNone -> anomaly (Pp.str "Cannot find equation lemma.")
in
- tclTHENSEQ[
+ tclTHENLIST[
tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) ids;
Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_lemma));
(* Don't forget to $\zeta$ normlize the term since the principles
@@ -719,7 +698,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
thin ids
]
else
- Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst f)))])
+ Proofview.V82.of_tactic (unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst (project g) f)))])
in
(* The proof of each branche itself *)
let ind_number = ref 0 in
@@ -739,7 +718,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
end
in
let this_branche_ids = List.nth intro_pats (pred i) in
- tclTHENSEQ[
+ tclTHENLIST[
(* we expand the definition of the function *)
observe_tac "rewrite_tac" (rewrite_tac this_ind_number this_branche_ids);
(* introduce hypothesis with some rewrite *)
@@ -750,8 +729,9 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
g
in
let params_names = fst (List.chop princ_infos.nparams args_names) in
+ let open EConstr in
let params = List.map mkVar params_names in
- tclTHENSEQ
+ tclTHENLIST
[ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) (args_names@[res;hres]);
observe_tac "h_generalize"
(Proofview.V82.of_tactic (generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)]));
@@ -763,19 +743,20 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
g
-(* [derive_correctness make_scheme functional_induction funs graphs] create correctness and completeness
+(* [derive_correctness make_scheme 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: pconstant list) (graphs:inductive list) =
+let derive_correctness make_scheme (funs: pconstant list) (graphs:inductive list) =
assert (funs <> []);
assert (graphs <> []);
let funs = Array.of_list funs and graphs = Array.of_list graphs in
- let funs_constr = Array.map mkConstU funs in
- States.with_state_protection_on_exception
+ let map (c, u) = mkConstU (c, EInstance.make u) in
+ let funs_constr = Array.map map funs in
+ (* XXX STATE Why do we need this... why is the toplevel protection not enought *)
+ funind_purify
(fun () ->
let env = Global.env () in
let evd = ref (Evd.from_env env) in
@@ -789,10 +770,10 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
in
let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in
graphs_constr.(i) <- graph;
- let type_of_lemma = Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in
+ let type_of_lemma = EConstr.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt in
let _ = Typing.e_type_of (Global.env ()) evd type_of_lemma in
let type_of_lemma = nf_zeta type_of_lemma in
- observe (str "type_of_lemma := " ++ Printer.pr_lconstr_env (Global.env ()) !evd type_of_lemma);
+ observe (str "type_of_lemma := " ++ Printer.pr_leconstr_env (Global.env ()) !evd type_of_lemma);
type_of_lemma,type_info
)
funs_constr
@@ -811,18 +792,18 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
Array.of_list
(List.map
(fun entry ->
- (fst (fst(Future.force entry.Entries.const_entry_body)), Option.get entry.Entries.const_entry_type )
+ (EConstr.of_constr (fst (fst(Future.force entry.Entries.const_entry_body))), EConstr.of_constr (Option.get entry.Entries.const_entry_type ))
)
- (make_scheme evd (Array.map_to_list (fun const -> const,GType []) funs))
+ (make_scheme evd (Array.map_to_list (fun const -> const,Sorts.InType) funs))
)
)
in
let proving_tac =
- prove_fun_correct !evd functional_induction funs_constr graphs_constr schemes lemmas_types_infos
+ prove_fun_correct !evd funs_constr graphs_constr schemes lemmas_types_infos
in
Array.iteri
(fun i f_as_constant ->
- let f_id = Label.to_id (con_label (fst f_as_constant)) in
+ let f_id = Label.to_id (Constant.label (fst f_as_constant)) in
(*i The next call to mk_correct_id is valid since we are constructing the lemma
Ensures by: obvious
i*)
@@ -842,7 +823,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
(* let lem_cst = fst (destConst (Constrintern.global_reference lem_id)) in *)
let _,lem_cst_constr = Evd.fresh_global
(Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in
- let (lem_cst,_) = destConst lem_cst_constr in
+ let lem_cst_constr = EConstr.of_constr lem_cst_constr in
+ let (lem_cst,_) = destConst !evd lem_cst_constr in
update_Function {finfo with correctness_lemma = Some lem_cst};
)
@@ -856,23 +838,23 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
let type_info = (type_of_lemma_ctxt,type_of_lemma_concl) in
graphs_constr.(i) <- graph;
let type_of_lemma =
- Termops.it_mkProd_or_LetIn type_of_lemma_concl type_of_lemma_ctxt
+ EConstr.it_mkProd_or_LetIn 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);
+ observe (str "type_of_lemma := " ++ Printer.pr_leconstr_env env !evd type_of_lemma);
type_of_lemma,type_info
)
funs_constr
graphs_constr
in
- let (kn,_) as graph_ind,u = (destInd graphs_constr.(0)) in
+ let (kn,_) as graph_ind,u = (destInd !evd graphs_constr.(0)) in
let mib,mip = Global.lookup_inductive graph_ind in
let sigma, scheme =
(Indrec.build_mutual_induction_scheme (Global.env ()) !evd
(Array.to_list
(Array.mapi
- (fun i _ -> ((kn,i),u(* Univ.Instance.empty *)),true,InType)
+ (fun i _ -> ((kn,i), EInstance.kind !evd u),true,InType)
mib.Declarations.mind_packets
)
)
@@ -886,7 +868,7 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
in
Array.iteri
(fun i f_as_constant ->
- let f_id = Label.to_id (con_label (fst f_as_constant)) in
+ let f_id = Label.to_id (Constant.label (fst f_as_constant)) in
(*i The next call to mk_complete_id is valid since we are constructing the lemma
Ensures by: obvious
i*)
@@ -902,7 +884,8 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
let finfo = find_Function_infos (fst f_as_constant) in
let _,lem_cst_constr = Evd.fresh_global
(Global.env ()) !evd (Constrintern.locate_reference (Libnames.qualid_of_ident lem_id)) in
- let (lem_cst,_) = destConst lem_cst_constr in
+ let lem_cst_constr = EConstr.of_constr lem_cst_constr in
+ let (lem_cst,_) = destConst !evd lem_cst_constr in
update_Function {finfo with completeness_lemma = Some lem_cst}
)
funs)
@@ -917,16 +900,17 @@ let derive_correctness make_scheme functional_induction (funs: pconstant list) (
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 sigma = project g in
let typ = pf_unsafe_type_of g (mkVar hid) in
- match kind_of_term typ with
- | App(i,args) when isInd i ->
- let ((kn',num) as ind'),u = destInd i in
+ match EConstr.kind sigma typ with
+ | App(i,args) when isInd sigma i ->
+ let ((kn',num) as ind'),u = destInd sigma i in
if MutInd.equal 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 (Pp.str "Cannot retrieve infos about a mutual block")
+ anomaly (Pp.str "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
@@ -935,7 +919,7 @@ let revert_graph kn post_tac hid g =
| None -> tclIDTAC g
| Some f_complete ->
let f_args,res = Array.chop (Array.length args - 1) args in
- tclTHENSEQ
+ tclTHENLIST
[
Proofview.V82.of_tactic (generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])]);
thin [hid];
@@ -965,21 +949,22 @@ let revert_graph kn post_tac hid g =
\end{enumerate}
*)
-let functional_inversion kn hid fconst f_correct : tactic =
+let functional_inversion kn hid fconst f_correct : Tacmach.tactic =
fun g ->
let old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty in
+ let sigma = project g in
let type_of_h = pf_unsafe_type_of g (mkVar hid) in
- match kind_of_term type_of_h with
- | App(eq,args) when eq_constr eq (make_eq ()) ->
+ match EConstr.kind sigma type_of_h with
+ | App(eq,args) when EConstr.eq_constr sigma eq (make_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 ->
+ match EConstr.kind sigma args.(1),EConstr.kind sigma args.(2) with
+ | App(f,f_args),_ when EConstr.eq_constr sigma f fconst ->
((fun hid -> Proofview.V82.of_tactic (intros_symmetry (Locusops.onHyp hid))),f_args,args.(2))
- |_,App(f,f_args) when eq_constr f fconst ->
+ |_,App(f,f_args) when EConstr.eq_constr sigma f fconst ->
((fun hid -> tclIDTAC),f_args,args.(1))
| _ -> (fun hid -> tclFAIL 1 (mt ())),[||],args.(2)
in
- tclTHENSEQ[
+ tclTHENLIST [
pre_tac hid;
Proofview.V82.of_tactic (generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])]);
thin [hid];
@@ -993,12 +978,13 @@ let functional_inversion kn hid fconst f_correct : tactic =
| _ -> tclFAIL 1 (mt ()) g
+let error msg = user_err Pp.(str msg)
let invfun qhyp f =
let f =
match f with
| ConstRef f -> f
- | _ -> raise (CErrors.UserError("",str "Not a function"))
+ | _ -> raise (CErrors.UserError(None,str "Not a function"))
in
try
let finfos = find_Function_infos f in
@@ -1012,7 +998,7 @@ let invfun qhyp f =
| Not_found -> error "No graph found"
| Option.IsNone -> error "Cannot use equivalence with graph!"
-
+exception NoFunction
let invfun qhyp f g =
match f with
| Some f -> invfun qhyp f g
@@ -1020,42 +1006,43 @@ let invfun qhyp f g =
Proofview.V82.of_tactic begin
Tactics.try_intros_until
(fun hid -> Proofview.V82.tactic begin fun g ->
+ let sigma = project g in
let hyp_typ = pf_unsafe_type_of g (mkVar hid) in
- match kind_of_term hyp_typ with
- | App(eq,args) when eq_constr eq (make_eq ()) ->
+ match EConstr.kind sigma hyp_typ with
+ | App(eq,args) when EConstr.eq_constr sigma eq (make_eq ()) ->
begin
- let f1,_ = decompose_app args.(1) in
+ let f1,_ = decompose_app sigma args.(1) in
try
- if not (isConst f1) then failwith "";
- let finfos = find_Function_infos (fst (destConst f1)) in
+ if not (isConst sigma f1) then raise NoFunction;
+ let finfos = find_Function_infos (fst (destConst sigma 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 ->
+ with | NoFunction | Option.IsNone | Not_found ->
try
- let f2,_ = decompose_app args.(2) in
- if not (isConst f2) then failwith "";
- let finfos = find_Function_infos (fst (destConst f2)) in
+ let f2,_ = decompose_app sigma args.(2) in
+ if not (isConst sigma f2) then raise NoFunction;
+ let finfos = find_Function_infos (fst (destConst sigma 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 least one Function")
+ | NoFunction ->
+ user_err (str "Hypothesis " ++ Ppconstr.pr_id hid ++ str " must contain at least 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)
+ else user_err (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)
+ else user_err (str "Cannot find inversion information for hypothesis " ++ Ppconstr.pr_id hid)
end
- | _ -> errorlabstrm "" (Ppconstr.pr_id hid ++ str " must be an equality ")
+ | _ -> user_err (Ppconstr.pr_id hid ++ str " must be an equality ")
end)
qhyp
end