summaryrefslogtreecommitdiff
path: root/plugins/funind
diff options
context:
space:
mode:
authorGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
committerGravatar Enrico Tassi <gareuselesinge@debian.org>2015-01-25 14:42:51 +0100
commit7cfc4e5146be5666419451bdd516f1f3f264d24a (patch)
treee4197645da03dc3c7cc84e434cc31d0a0cca7056 /plugins/funind
parent420f78b2caeaaddc6fe484565b2d0e49c66888e5 (diff)
Imported Upstream version 8.5~beta1+dfsg
Diffstat (limited to 'plugins/funind')
-rw-r--r--plugins/funind/Recdef.v38
-rw-r--r--plugins/funind/functional_principles_proofs.ml474
-rw-r--r--plugins/funind/functional_principles_types.ml228
-rw-r--r--plugins/funind/functional_principles_types.mli10
-rw-r--r--plugins/funind/g_indfun.ml4167
-rw-r--r--plugins/funind/glob_term_to_relation.ml399
-rw-r--r--plugins/funind/glob_term_to_relation.mli10
-rw-r--r--plugins/funind/glob_termops.ml168
-rw-r--r--plugins/funind/glob_termops.mli58
-rw-r--r--plugins/funind/indfun.ml470
-rw-r--r--plugins/funind/indfun.mli15
-rw-r--r--plugins/funind/indfun_common.ml212
-rw-r--r--plugins/funind/indfun_common.mli52
-rw-r--r--plugins/funind/invfun.ml706
-rw-r--r--plugins/funind/merge.ml195
-rw-r--r--plugins/funind/recdef.ml2077
-rw-r--r--plugins/funind/recdef.mli20
17 files changed, 2741 insertions, 2558 deletions
diff --git a/plugins/funind/Recdef.v b/plugins/funind/Recdef.v
index 51ede26e..a63941f0 100644
--- a/plugins/funind/Recdef.v
+++ b/plugins/funind/Recdef.v
@@ -1,10 +1,13 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
+
+Require Import PeanoNat.
+
Require Compare_dec.
Require Wf_nat.
@@ -19,30 +22,29 @@ Fixpoint iter (n : nat) : (A -> A) -> A -> A :=
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.
+Theorem le_lt_SS x y : x <= y -> x < S (S y).
+Proof.
+ intros. now apply Nat.lt_succ_r, Nat.le_le_succ_r.
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.
+Theorem Splus_lt x y : y < S (x + y).
+Proof.
+ apply Nat.lt_succ_r. rewrite Nat.add_comm. apply Nat.le_add_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.
+Theorem SSplus_lt x y : x < S (S (x + y)).
+Proof.
+ apply le_lt_SS, Nat.le_add_r.
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].
+Definition max m n : max_type m n.
+Proof.
+ destruct (Compare_dec.le_gt_dec m n) as [h|h].
+ - exists n; [exact h | apply le_n].
+ - exists m; [apply le_n | apply Nat.lt_le_incl; exact h].
Defined.
+
+Definition Acc_intro_generator_function := fun A R => @Acc_intro_generator A R 100.
diff --git a/plugins/funind/functional_principles_proofs.ml b/plugins/funind/functional_principles_proofs.ml
index b5876ffa..c8214ada 100644
--- a/plugins/funind/functional_principles_proofs.ml
+++ b/plugins/funind/functional_principles_proofs.ml
@@ -1,45 +1,37 @@
open Printer
+open Errors
open Util
open Term
+open Vars
+open Context
open Namegen
open Names
open Declarations
+open Declareops
open Pp
-open Entries
-open Hiddentac
-open Evd
open Tacmach
open Proof_type
open Tacticals
open Tactics
open Indfun_common
open Libnames
+open Globnames
+open Misctypes
-let msgnl = Pp.msgnl
-
+(* 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
+ then Pp.msg_debug strm
else ()
-
-
-
let do_observe_tac s tac g =
try let v = tac g in (* msgnl (goal ++ fnl () ++ (str s)++(str " ")++(str "finished")); *) v
- with reraise ->
- let e = Cerrors.process_vernac_interp_error reraise in
- let goal =
- try (Printer.pr_goal g)
- with e when Errors.noncritical e -> assert false
- in
- msgnl (str "observation "++ s++str " raised exception " ++
+ with e ->
+ let e = Cerrors.process_vernac_interp_error e in
+ let goal = begin try (Printer.pr_goal g) with _ -> assert false end in
+ msg_debug (str "observation "++ s++str " raised exception " ++
Errors.print e ++ str " on goal " ++ goal );
raise e;;
@@ -49,16 +41,55 @@ let observe_tac_stream s tac g =
else tac g
let observe_tac s tac g = observe_tac_stream (str s) tac g
+ *)
+
+
+let debug_queue = Stack.create ()
+
+let rec print_debug_queue b e =
+ if not (Stack.is_empty debug_queue)
+ then
+ begin
+ let lmsg,goal = Stack.pop debug_queue in
+ if b then
+ Pp.msg_debug (lmsg ++ (str " raised exception " ++ Errors.print e) ++ str " on goal " ++ goal)
+ else
+ begin
+ Pp.msg_debug (str " from " ++ lmsg ++ str " on goal " ++ goal);
+ end;
+ print_debug_queue false e;
+ end
-(* let tclTRYD tac = *)
-(* if !Flags.debug || do_observe () *)
-(* then (fun g -> try (\* do_observe_tac "" *\)tac g with _ -> tclIDTAC g) *)
-(* else tac *)
+let observe strm =
+ if do_observe ()
+ then Pp.msg_debug strm
+ else ()
+
+let do_observe_tac s tac g =
+ let goal = Printer.pr_goal g in
+ let lmsg = (str "observation : ") ++ s in
+ Stack.push (lmsg,goal) debug_queue;
+ try
+ let v = tac g in
+ ignore(Stack.pop debug_queue);
+ v
+ with reraise ->
+ let reraise = Errors.push reraise in
+ if not (Stack.is_empty debug_queue)
+ then print_debug_queue true (fst (Cerrors.process_vernac_interp_error reraise));
+ iraise reraise
+
+let observe_tac_stream s tac g =
+ if do_observe ()
+ then do_observe_tac s tac g
+ else tac g
+let observe_tac s = observe_tac_stream (str s)
+
let list_chop ?(msg="") n l =
try
- list_chop n l
+ List.chop n l
with Failure (msg') ->
failwith (msg ^ msg')
@@ -70,17 +101,17 @@ let make_refl_eq constructor type_of_t t =
type pte_info =
{
- proving_tac : (identifier list -> Tacmach.tactic);
+ proving_tac : (Id.t list -> Tacmach.tactic);
is_valid : constr -> bool
}
-type ptes_info = pte_info Idmap.t
+type ptes_info = pte_info Id.Map.t
type 'a dynamic_info =
{
nb_rec_hyps : int;
- rec_hyps : identifier list ;
- eq_hyps : identifier list;
+ rec_hyps : Id.t list ;
+ eq_hyps : Id.t list;
info : 'a
}
@@ -89,28 +120,17 @@ type body_info = constr dynamic_info
let finish_proof dynamic_infos g =
observe_tac "finish"
- ( h_assumption)
+ (Proofview.V82.of_tactic assumption)
g
let refine c =
- Tacmach.refine_no_check c
+ Tacmach.refine 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 eq_constr u v = eq_constr_nounivs u v
let is_trivial_eq t =
let res = try
@@ -157,11 +177,11 @@ 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)))
+ ((* observe_tac msg *) Proofview.V82.of_tactic (assert_by (Name prov_id) t (Proofview.V82.tactic (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])
+ (* observe_tac "change_hyp_with_using rename " *) (Proofview.V82.of_tactic (rename_hyp [prov_id,hyp_id]))
]] g
exception TOREMOVE
@@ -171,7 +191,7 @@ 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 *)
+ tclDO nb_intros (Proofview.V82.of_tactic intro); (* introducing context *)
(fun g ->
let context_hyps =
fst (list_chop ~msg:"prove_trivial_eq : " nb_intros (pf_ids_of_hyps g))
@@ -188,7 +208,7 @@ let prove_trivial_eq h_id context (constructor,type_of_term,term) =
let find_rectype env c =
- let (t, l) = decompose_app (Reduction.whd_betaiotazeta c) in
+ let (t, l) = decompose_app (Reduction.whd_betaiotazeta env c) in
match kind_of_term t with
| Ind ind -> (t, l)
| Construct _ -> (t,l)
@@ -216,7 +236,7 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
failwith "NoChange";
end
in
- let eq_constr = Reductionops.is_conv env sigma in
+ let eq_constr = Evarconv.e_conv env (ref 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";
@@ -245,12 +265,12 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
let t2 = destRel t2 in
begin
try
- let t1' = Intmap.find t2 sub in
+ let t1' = Int.Map.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
+ Int.Map.add t2 t1 sub
end
else if isAppConstruct t1 && isAppConstruct t2
then
@@ -264,18 +284,17 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
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 Int.Map.empty (snd t1) (snd t2) in
let sub = compute_substitution sub (fst t1) (fst t2) in
let end_of_type_with_pop = Termops.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
+ let sub = Int.Map.bindings 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''
+ sub
in
let old_context_length = List.length context + 1 in
let witness_fun =
@@ -284,11 +303,11 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
)
in
let new_type_of_hyp,ctxt_size,witness_fun =
- list_fold_left_i
+ 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!";
+ let witness = Int.Map.find i sub in
+ if not (Option.is_empty b') then anomaly (Pp.str "can not redefine a rel!");
(Termops.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)
@@ -304,12 +323,13 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
in
let prove_new_hyp : tactic =
tclTHEN
- (tclDO ctxt_size intro)
+ (tclDO ctxt_size (Proofview.V82.of_tactic 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
+ let evm, _ = pf_apply Typing.e_type_of g to_refine in
+ tclTHEN (Refiner.tclEVARS evm) (refine to_refine) g
)
in
let simpl_eq_tac =
@@ -332,14 +352,14 @@ let change_eq env sigma hyp_id (context:rel_context) x t end_of_type =
new_ctxt,new_end_of_type,simpl_eq_tac
-let is_property ptes_info t_x full_type_of_hyp =
+let is_property (ptes_info: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
+ if isVar pte && Array.for_all closed0 args
then
try
- let info = Idmap.find (destVar pte) ptes_info in
+ let info = Id.Map.find (destVar pte) ptes_info in
info.is_valid full_type_of_hyp
with Not_found -> false
else false
@@ -352,10 +372,10 @@ let isLetIn t =
let h_reduce_with_zeta =
- h_reduce
- (Glob_term.Cbv
- {Glob_term.all_flags
- with Glob_term.rDelta = false;
+ reduce
+ (Genredexpr.Cbv
+ {Redops.all_flags
+ with Genredexpr.rDelta = false;
})
@@ -374,17 +394,17 @@ let rewrite_until_var arg_num eq_ids : tactic =
then tclIDTAC g
else
match eq_ids with
- | [] -> anomaly "Cannot find a way to prove recursive property";
+ | [] -> anomaly (Pp.str "Cannot find a way to prove recursive property");
| eq_id::eq_ids ->
tclTHEN
- (tclTRY (Equality.rewriteRL (mkVar eq_id)))
+ (tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar eq_id))))
(do_rewrite eq_ids)
g
in
do_rewrite eq_ids
-let rec_pte_id = id_of_string "Hrec"
+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
@@ -398,13 +418,8 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
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
-
- ]
+ [ h_reduce_with_zeta (Locusops.onHyp hyp_id);
+ scan_type new_context new_typ_of_hyp ]
else if isProd type_of_hyp
then
begin
@@ -413,14 +428,14 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
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 (* fix_info *) prove_rec_hyp = (Id.Map.find (destVar pte) ptes_infos).proving_tac in
let popped_t' = Termops.pop t' in
let real_type_of_hyp = it_mkProd_or_LetIn popped_t' context in
let prove_new_type_of_hyp =
let context_length = List.length context in
tclTHENLIST
[
- tclDO context_length intro;
+ tclDO context_length (Proofview.V82.of_tactic intro);
(fun g ->
let context_hyps_ids =
fst (list_chop ~msg:"rec hyp : context_hyps"
@@ -434,7 +449,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
in
(* observe_tac "rec hyp " *)
(tclTHENS
- (assert_tac (Name rec_pte_id) t_x)
+ (Proofview.V82.of_tactic (assert_before (Name rec_pte_id) t_x))
[
(* observe_tac "prove rec hyp" *) (prove_rec_hyp eq_hyps);
(* observe_tac "prove rec hyp" *)
@@ -471,7 +486,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
let prove_trivial =
let nb_intro = List.length context in
tclTHENLIST [
- tclDO nb_intro intro;
+ tclDO nb_intro (Proofview.V82.of_tactic intro);
(fun g ->
let context_hyps =
fst (list_chop ~msg:"removing True : context_hyps "nb_intro (pf_ids_of_hyps g))
@@ -533,7 +548,7 @@ let clean_hyp_with_heq ptes_infos eq_hyps hyp_id env sigma =
thin [hyp_id],[]
-let clean_goal_with_heq ptes_infos continue_tac dyn_infos =
+let clean_goal_with_heq ptes_infos continue_tac (dyn_infos:body_info) =
fun g ->
let env = pf_env g
and sigma = project g
@@ -562,7 +577,7 @@ let clean_goal_with_heq ptes_infos continue_tac dyn_infos =
]
g
-let heq_id = id_of_string "Heq"
+let heq_id = Id.of_string "Heq"
let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos =
fun g ->
@@ -570,12 +585,12 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos =
tclTHENLIST
[
(* We first introduce the variables *)
- tclDO nb_first_intro (intro_avoiding dyn_infos.rec_hyps);
+ tclDO nb_first_intro (Proofview.V82.of_tactic (intro_avoiding dyn_infos.rec_hyps));
(* Then the equation itself *)
- intro_using heq_id;
+ Proofview.V82.of_tactic (intro_using heq_id);
onLastHypId (fun heq_id -> tclTHENLIST [
(* Then the new hypothesis *)
- tclMAP introduction_no_check dyn_infos.rec_hyps;
+ tclMAP (fun id -> Proofview.V82.of_tactic (introduction ~check:false id)) 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
@@ -585,9 +600,9 @@ let treat_new_case ptes_infos nb_prod continue_tac term dyn_infos =
| 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
+ pr_lconstr_env (pf_env g') Evd.empty new_term_value_eq
);
- anomaly "cannot compute new term value"
+ anomaly (Pp.str "cannot compute new term value")
in
let fun_body =
mkLambda(Anonymous,
@@ -615,17 +630,20 @@ let my_orelse tac1 tac2 g =
(* observe (str "using snd tac since : " ++ Errors.print e); *)
tac2 g
-let instanciate_hyps_with_args (do_prove:identifier list -> tactic) hyps args_id =
+let instanciate_hyps_with_args (do_prove:Id.t 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
+ let c = mkApp(mkVar hid,args) in
+ let evm, _ = pf_apply Typing.e_type_of g c in
tclTHENLIST[
- pose_proof (Name prov_hid) (mkApp(mkVar hid,args));
+ Refiner.tclEVARS evm;
+ Proofview.V82.of_tactic (pose_proof (Name prov_hid) c);
thin [hid];
- h_rename [prov_hid,hid]
+ Proofview.V82.of_tactic (rename_hyp [prov_hid,hid])
] g
)
( (*
@@ -642,23 +660,23 @@ let instanciate_hyps_with_args (do_prove:identifier list -> tactic) hyps args_id
)
)
in
- if args_id = []
+ if List.is_empty args_id
then
tclTHENLIST [
- tclMAP (fun hyp_id -> h_reduce_with_zeta (Tacticals.onHyp hyp_id)) hyps;
+ tclMAP (fun hyp_id -> h_reduce_with_zeta (Locusops.onHyp hyp_id)) hyps;
do_prove hyps
]
else
tclTHENLIST
[
- tclMAP (fun hyp_id -> h_reduce_with_zeta (Tacticals.onHyp hyp_id)) hyps;
+ tclMAP (fun hyp_id -> h_reduce_with_zeta (Locusops.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
+ List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.empty
in
let remaining_hyps =
- List.filter (fun id -> Idset.mem id all_g_hyps_id) hyps
+ List.filter (fun id -> Id.Set.mem id all_g_hyps_id) hyps
in
do_prove remaining_hyps g
)
@@ -687,11 +705,11 @@ let build_proof
in
tclTHENSEQ
[
- h_generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps));
+ Simple.generalize (term_eq::(List.map mkVar dyn_infos.rec_hyps));
thin dyn_infos.rec_hyps;
- pattern_option [(false,[1]),t] None;
+ pattern_option [Locus.AllOccurrencesBut [1],t] None;
(fun g -> observe_tac "toto" (
- tclTHENSEQ [h_simplest_case t;
+ tclTHENSEQ [Proofview.V82.of_tactic (Simple.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
@@ -716,7 +734,7 @@ let build_proof
match kind_of_term( pf_concl g) with
| Prod _ ->
tclTHEN
- intro
+ (Proofview.V82.of_tactic intro)
(fun g' ->
let (id,_,_) = pf_last_hyp g' in
let new_term =
@@ -746,6 +764,7 @@ let build_proof
begin
match kind_of_term f with
| App _ -> assert false (* we have collected all the app in decompose_app *)
+ | Proj _ -> assert false (*FIXME*)
| Var _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _ | Sort _ | Prod _ ->
let new_infos =
{ dyn_infos with
@@ -753,7 +772,7 @@ let build_proof
}
in
build_proof_args do_finalize new_infos g
- | Const c when not (List.mem c fnames) ->
+ | Const (c,_) when not (List.mem_f Constant.equal c fnames) ->
let new_infos =
{ dyn_infos with
info = (f,args)
@@ -775,9 +794,10 @@ let build_proof
tclTHENLIST
[tclMAP
- (fun hyp_id -> h_reduce_with_zeta (Tacticals.onHyp hyp_id))
+ (fun hyp_id ->
+ h_reduce_with_zeta (Locusops.onHyp hyp_id))
dyn_infos.rec_hyps;
- h_reduce_with_zeta Tacticals.onConcl;
+ h_reduce_with_zeta Locusops.onConcl;
build_proof do_finalize new_infos
]
g
@@ -797,6 +817,7 @@ let build_proof
| Fix _ | CoFix _ ->
error ( "Anonymous local (co)fixpoints are not handled yet")
+ | Proj _ -> error "Prod"
| Prod _ -> error "Prod"
| LetIn _ ->
let new_infos =
@@ -807,28 +828,28 @@ let build_proof
tclTHENLIST
[tclMAP
- (fun hyp_id -> h_reduce_with_zeta (Tacticals.onHyp hyp_id))
+ (fun hyp_id -> h_reduce_with_zeta (Locusops.onHyp hyp_id))
dyn_infos.rec_hyps;
- h_reduce_with_zeta Tacticals.onConcl;
+ h_reduce_with_zeta Locusops.onConcl;
build_proof do_finalize new_infos
] g
- | Rel _ -> anomaly "Free var in goal conclusion !"
+ | Rel _ -> anomaly (Pp.str "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
+ observe_tac_stream (str "build_proof with " ++ Printer.pr_lconstr dyn_infos.info ) (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
- | [] ->
+ 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) *)
-(* ); *)
+ | 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 *)
@@ -842,14 +863,14 @@ let build_proof
g
in
(* observe_tac "build_proof_args" *) (tac ) g
- in
- let do_finish_proof dyn_infos =
+ in
+ let do_finish_proof dyn_infos =
(* tclTRYD *) (clean_goal_with_heq
- ptes_infos
- finish_proof dyn_infos)
+ ptes_infos
+ finish_proof dyn_infos)
in
- (* observe_tac "build_proof" *)
- (build_proof (clean_goal_with_heq ptes_infos do_finish_proof) dyn_infos)
+ (* observe_tac "build_proof" *)
+ (build_proof (clean_goal_with_heq ptes_infos do_finish_proof) dyn_infos)
@@ -863,18 +884,11 @@ let build_proof
(* 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;
+ name : Id.t;
types : types;
offset : int;
nb_realargs : int;
@@ -901,9 +915,6 @@ let prove_rec_hyp 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
@@ -911,17 +922,17 @@ let generalize_non_dep hyp g =
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 (Termops.occur_var_in_decl env hyp) keep
- or Termops.occur_var env hyp hyp_typ
- or Termops.is_section_variable hyp (* should be dangerous *)
+ if Id.List.mem hyp hyps
+ || List.exists (Termops.occur_var_in_decl env hyp) keep
+ || Termops.occur_var env hyp hyp_typ
+ || 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 "h_generalize" *) (Simple.generalize (List.map mkVar to_revert) ))
((* observe_tac "thin" *) (thin to_revert))
g
@@ -936,11 +947,9 @@ 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 f_def = Global.lookup_constant (fst (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 (body_of_constant f_def))
- in
+ let f_body = Option.get (Global.body_of_constant_body f_def)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 =
@@ -955,20 +964,20 @@ let generate_equation_lemma fnames f fun_num nb_params nb_args rec_args_num =
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
+ (Typeops.type_of_constant_type (Global.env ()) (*FIXME*)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 eqn type_ctxt in
- let f_id = id_of_label (con_label (destConst f)) in
+ let f_id = Label.to_id (con_label (fst (destConst f))) in
let prove_replacement =
tclTHENSEQ
[
- tclDO (nb_params + rec_args_num + 1) intro;
+ tclDO (nb_params + rec_args_num + 1) (Proofview.V82.of_tactic 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,Glob_term.NoBindings));
- intros_reflexivity] g
+ (* observe_tac "h_case" *) (Proofview.V82.of_tactic (simplest_case (mkVar rec_id)));
+ (Proofview.V82.of_tactic intros_reflexivity)] g
)
]
in
@@ -977,11 +986,12 @@ let generate_equation_lemma fnames f fun_num nb_params nb_args rec_args_num =
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
+ (Decl_kinds.Global, false, (Decl_kinds.Proof Decl_kinds.Theorem))
+ Evd.empty
+ lemma_type
+ (Lemmas.mk_hook (fun _ _ -> ()));
+ ignore (Pfedit.by (Proofview.V82.tactic prove_replacement));
+ Lemmas.save_proof (Vernacexpr.Proved(false,None))
@@ -989,10 +999,10 @@ let generate_equation_lemma fnames f fun_num nb_params nb_args rec_args_num =
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
+ let finfos = find_Function_infos (fst (destConst f)) (*FIXME*) 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
+ let f_id = Label.to_id (con_label (fst (destConst f))) in
(*i The next call to mk_equation_id is valid since we will construct the lemma
Ensures by: obvious
i*)
@@ -1001,12 +1011,12 @@ let do_replace params rec_arg_num rev_args_id f fun_num all_funs g =
let _ =
match e with
| Option.IsNone ->
- let finfos = find_Function_infos (destConst f) in
+ let finfos = find_Function_infos (fst (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"
+ | _ -> Errors.anomaly (Pp.str "Not a constant")
)
}
| _ -> ()
@@ -1016,12 +1026,12 @@ let do_replace params rec_arg_num rev_args_id f fun_num all_funs g =
in
let nb_intro_to_do = nb_prod (pf_concl g) in
tclTHEN
- (tclDO nb_intro_to_do intro)
+ (tclDO nb_intro_to_do (Proofview.V82.of_tactic 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'
+ tclTHEN (Proofview.V82.of_tactic (Equality.rewriteLR equation_lemma)) (revert just_introduced_id) g'
)
g
@@ -1034,7 +1044,7 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
(fun na ->
let new_id =
match na with
- Name id -> fresh_id !avoid (string_of_id id)
+ Name id -> fresh_id !avoid (Id.to_string id)
| Anonymous -> fresh_id !avoid "H"
in
avoid := new_id :: !avoid;
@@ -1055,9 +1065,8 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
}
in
let get_body const =
- match body_of_constant (Global.lookup_constant const) with
- | Some b ->
- let body = force b in
+ match Global.body_of_constant const with
+ | Some body ->
Tacred.cbv_norm_flags
(Closure.RedFlags.mkflags [Closure.RedFlags.fZETA])
(Global.env ())
@@ -1137,7 +1146,7 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
typess
in
let pte_to_fix,rev_info =
- list_fold_left_i
+ 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
@@ -1175,14 +1184,14 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
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)
+ (Id.Map.add (Nameops.out_name pte) info acc_map,info::acc_info)
)
0
- (Idmap.empty,[])
+ (Id.Map.empty,[])
(List.rev princ_info.predicates)
in
pte_to_fix,List.rev rev_info
- | _ -> Idmap.empty,[]
+ | _ -> Id.Map.empty,[]
in
let mk_fixes : tactic =
let pre_info,infos = list_chop fun_num infos in
@@ -1194,19 +1203,19 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
(fun fi -> fi.name,fi.idx + 1 ,fi.types)
(pre_info@others_infos)
in
- if other_fix_infos = []
+ if List.is_empty other_fix_infos
then
- (* observe_tac ("h_fix") *) (h_fix (Some this_fix_info.name) (this_fix_info.idx +1))
+ (* observe_tac ("h_fix") *) (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"
+ Tactics.mutual_fix this_fix_info.name (this_fix_info.idx + 1)
+ other_fix_infos 0
+ | _ -> anomaly (Pp.str "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 "introducing params" *) Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.params));
+ (* observe_tac "introducing predictes" *) Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.predicates));
+ (* observe_tac "introducing branches" *) Proofview.V82.of_tactic (intros_using (List.rev_map id_of_decl princ_info.branches));
(* observe_tac "building fixes" *) mk_fixes;
]
in
@@ -1217,14 +1226,13 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
try
let pte =
try destVar pte
- with e when Errors.noncritical e ->
- anomaly "Property is not a variable"
+ with DestKO -> anomaly (Pp.str "Property is not a variable")
in
- let fix_info = Idmap.find pte ptes_to_fix in
+ let fix_info = Id.Map.find pte ptes_to_fix in
let nb_args = fix_info.nb_realargs in
tclTHENSEQ
[
- (* observe_tac ("introducing args") *) (tclDO nb_args intro);
+ (* observe_tac ("introducing args") *) (tclDO nb_args (Proofview.V82.of_tactic 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
@@ -1258,7 +1266,7 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
build_proof
interactive_proof
(Array.to_list fnames)
- (Idmap.map prove_rec_hyp ptes_to_fix)
+ (Id.Map.map prove_rec_hyp ptes_to_fix)
in
let prove_tac branches =
let dyn_infos =
@@ -1268,7 +1276,7 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
}
in
observe_tac "cleaning" (clean_goal_with_heq
- (Idmap.map prove_rec_hyp ptes_to_fix)
+ (Id.Map.map prove_rec_hyp ptes_to_fix)
do_prove
dyn_infos)
in
@@ -1288,7 +1296,7 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
let nb_args = min (princ_info.nargs) (List.length ctxt) in
tclTHENSEQ
[
- tclDO nb_args intro;
+ tclDO nb_args (Proofview.V82.of_tactic 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
@@ -1307,12 +1315,12 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
in
let fname = destConst (fst (decompose_app (List.hd (List.rev pte_args)))) in
tclTHENSEQ
- [unfold_in_concl [(Termops.all_occurrences, Names.EvalConstRef fname)];
+ [unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst fname))];
let do_prove =
build_proof
interactive_proof
(Array.to_list fnames)
- (Idmap.map prove_rec_hyp ptes_to_fix)
+ (Id.Map.map prove_rec_hyp ptes_to_fix)
in
let prove_tac branches =
let dyn_infos =
@@ -1322,7 +1330,7 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
}
in
clean_goal_with_heq
- (Idmap.map prove_rec_hyp ptes_to_fix)
+ (Id.Map.map prove_rec_hyp ptes_to_fix)
do_prove
dyn_infos
in
@@ -1346,15 +1354,13 @@ let prove_princ_for_struct interactive_proof fun_num fnames all_funs _nparams :
(* 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 h_intros = Recdef.h_intros
-and list_rewrite = Recdef.list_rewrite
-and evaluable_of_global_reference = Recdef.evaluable_of_global_reference
+(* let 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 list_rewrite = Recdef.list_rewrite *)
+(* and evaluable_of_global_reference = Recdef.evaluable_of_global_reference *)
@@ -1362,7 +1368,7 @@ 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 !!"
+ | None -> anomaly (Pp.str "No tcc proof !!")
| Some lemma ->
fun gls ->
(* let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in *)
@@ -1387,14 +1393,14 @@ 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 )
+ tclFIRST (List.map (fun x -> Proofview.V82.of_tactic (Equality.rewriteRL x)) 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_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
+ 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
@@ -1402,17 +1408,6 @@ let backtrack_eqs_until_hrec hrec eqs : tactic =
backtrack gls
-
-let build_clause eqs =
- {
- Tacexpr.onhyps =
- Some (List.map
- (fun id -> (Glob_term.all_occurrences_expr, id), Termops.InHyp)
- eqs
- );
- Tacexpr.concl_occs = Glob_term.no_occurrences_expr
- }
-
let rec rewrite_eqs_in_eqs eqs =
match eqs with
| [] -> tclIDTAC
@@ -1422,8 +1417,9 @@ let rec rewrite_eqs_in_eqs eqs =
(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 Termops.all_occurrences true (* dep proofs also: *) true id (mkVar eq) false))
+ (Format.sprintf "rewrite %s in %s " (Id.to_string eq) (Id.to_string id))
+ (tclTRY (Proofview.V82.of_tactic (Equality.general_rewrite_in true Locus.AllOccurrences
+ true (* dep proofs also: *) true id (mkVar eq) false)))
gl
)
eqs
@@ -1435,22 +1431,22 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic =
(tclTHENSEQ
[
backtrack_eqs_until_hrec hrec eqs;
- (* observe_tac ("new_prove_with_tcc ( applying "^(string_of_id hrec)^" )" ) *)
+ (* observe_tac ("new_prove_with_tcc ( applying "^(Id.to_string hrec)^" )" ) *)
(tclTHENS (* We must have exactly ONE subgoal !*)
- (apply (mkVar hrec))
+ (Proofview.V82.of_tactic (apply (mkVar hrec)))
[ tclTHENSEQ
[
- keep (tcc_hyps@eqs);
- apply (Lazy.force acc_inv);
+ (Proofview.V82.of_tactic (keep (tcc_hyps@eqs)));
+ (Proofview.V82.of_tactic (apply (Lazy.force acc_inv)));
(fun g ->
if is_mes
then
- unfold_in_concl [(Termops.all_occurrences, evaluable_of_global_reference (delayed_force ltof_ref))] g
+ unfold_in_concl [(Locus.AllOccurrences, 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));
+ [tclTRY(list_rewrite false (List.map (fun v -> (mkVar v,true)) eqs));
observe_tac "rewrite_eqs_in_eqs" (rewrite_eqs_in_eqs eqs);
(observe_tac "finishing using"
(
@@ -1458,7 +1454,7 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic =
Eauto.eauto_with_bases
(true,5)
[Evd.empty,Lazy.force refl_equal]
- [Auto.Hint_db.empty empty_transparent_state false]
+ [Hints.Hint_db.empty empty_transparent_state false]
)
)
)
@@ -1471,13 +1467,13 @@ let new_prove_with_tcc is_mes acc_inv hrec tcc_hyps eqs : tactic =
let is_valid_hypothesis predicates_name =
- let predicates_name = List.fold_right Idset.add predicates_name Idset.empty in
+ let predicates_name = List.fold_right Id.Set.add predicates_name Id.Set.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
+ then Id.Set.mem (destVar pte) predicates_name
else false
else false
in
@@ -1499,7 +1495,7 @@ let prove_principle_for_gen
fun na ->
let new_id =
match na with
- | Name id -> fresh_id !avoid (string_of_id id)
+ | Name id -> fresh_id !avoid (Id.to_string id)
| Anonymous -> fresh_id !avoid "H"
in
avoid := new_id :: !avoid;
@@ -1531,7 +1527,7 @@ let prove_principle_for_gen
(* 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
+ Util.List.chop npost_rec_arg princ_info.args
in
let rec_arg_id =
match List.rev post_rec_arg with
@@ -1542,25 +1538,25 @@ let prove_principle_for_gen
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 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)))))
+ Nameops.out_name (fresh_id (Name (Id.of_string ("Acc_"^(Id.to_string rec_arg_id)))))
in
let revert l =
- tclTHEN (h_generalize (List.map mkVar l)) (clear l)
+ tclTHEN (Tactics.Simple.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)
+ (Proofview.V82.of_tactic (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))
+ (Proofview.V82.tactic (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|]))
+ Proofview.V82.of_tactic (Tactics.Simple.apply (mkApp(mkVar wf_thm_id,[|mkVar rec_arg_id|])))
)
)
)
@@ -1570,7 +1566,7 @@ let prove_principle_for_gen
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 !!")
+ | None -> error "No tcc proof !!"
| Some lemma -> lemma
in
(* let rec list_diff del_list check_list = *)
@@ -1588,18 +1584,18 @@ let prove_principle_for_gen
let hyps = pf_ids_of_hyps gls in
let hid =
next_ident_away_in_goal
- (id_of_string "prov")
+ (Id.of_string "prov")
hyps
in
tclTHENSEQ
[
generalize [lemma];
- h_intro hid;
- Elim.h_decompose_and (mkVar hid);
+ Proofview.V82.of_tactic (Simple.intro hid);
+ Proofview.V82.of_tactic (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 = []
+ tcc_list := List.rev (List.subtract Id.equal new_hyps (hid::hyps));
+ if List.is_empty !tcc_list
then
begin
tcc_list := [hid];
@@ -1617,22 +1613,22 @@ let prove_principle_for_gen
(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
+ (* observe_tac "" *) Proofview.V82.of_tactic (assert_by
(Name acc_rec_arg_id)
(mkApp (delayed_force acc_rel,[|input_type;relation;mkVar rec_arg_id|]))
- (prove_rec_arg_acc)
+ (Proofview.V82.tactic 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));
+ (* observe_tac "h_fix " *) (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);
+ Proofview.V82.of_tactic (Equality.rewriteLR (mkConst eq_ref));
(* observe_tac "finish" *) (fun gl' ->
let body =
let _,args = destApp (pf_concl gl') in
- array_last args
+ Array.last args
in
let body_info rec_hyps =
{
@@ -1677,14 +1673,14 @@ let prove_principle_for_gen
is_valid = is_valid_hypothesis predicates_names
}
in
- let ptes_info : pte_info Idmap.t =
+ let ptes_info : pte_info Id.Map.t =
List.fold_left
(fun map pte_id ->
- Idmap.add pte_id
+ Id.Map.add pte_id
pte_info
map
)
- Idmap.empty
+ Id.Map.empty
predicates_names
in
let make_proof rec_hyps =
diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml
index 04fcc8d4..545f8931 100644
--- a/plugins/funind/functional_principles_types.ml
+++ b/plugins/funind/functional_principles_types.ml
@@ -1,58 +1,25 @@
open Printer
+open Errors
open Util
open Term
+open Vars
+open Context
open Namegen
open Names
-open Declarations
+open Declareops
open Pp
open Entries
-open Hiddentac
-open Evd
-open Tacmach
-open Proof_type
-open Tacticals
open Tactics
open Indfun_common
open Functional_principles_proofs
+open Misctypes
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
+ then Pp.msg_debug s
(*
Transform an inductive induction principle into
@@ -63,14 +30,14 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
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 =
+ let rec change_predicates_names (avoid:Id.t 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 "
+ | (Anonymous,_,_)::_ -> anomaly (Pp.str "Anonymous property binder ")
in
let avoid = (Termops.ids_of_context env_with_params ) in
let princ_type_info =
@@ -91,7 +58,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
Nameops.out_name x,None,compose_prod real_args (mkSort new_sort)
in
let new_predicates =
- list_map_i
+ List.map_i
change_predicate_sort
0
princ_type_info.predicates
@@ -99,16 +66,16 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
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
+ | Some (Globnames.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
+ let set = List.fold_right Id.Set.add ptes_vars Id.Set.empty in
fun t ->
match kind_of_term t with
- | Var id -> Idset.mem id set
+ | Var id -> Id.Set.mem id set
| _ -> false
in
let pre_princ =
@@ -126,17 +93,17 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
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
+ | Ind((u,_),_) -> MutInd.equal u rel_as_kn
+ | Construct(((u,_),_),_) -> MutInd.equal u rel_as_kn
| _ -> false
in
let get_fun_num c =
match kind_of_term c with
- | Ind(_,num) -> num
- | Construct((_,num),_) -> num
+ | Ind((_,num),_) -> num
+ | Construct(((_,num),_),_) -> num
| _ -> assert false
in
- let dummy_var = mkVar (id_of_string "________") in
+ let dummy_var = mkVar (Id.of_string "________") in
let mk_replacement c i args =
let res = mkApp(rel_to_fun.(i), Array.map Termops.pop (array_get_start args)) in
(* observe (str "replacing " ++ pr_lconstr c ++ str " by " ++ pr_lconstr res); *)
@@ -157,7 +124,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
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 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) ->
@@ -191,7 +158,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
begin
try
let new_t,binders_to_remove_from_t = compute_new_princ_type remove env t in
- let new_x : name = get_name (Termops.ids_of_context env) x in
+ let new_x : Name.t = get_name (Termops.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
@@ -220,7 +187,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
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 (Termops.ids_of_context env) x in
+ let new_x : Name.t = get_name (Termops.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
@@ -255,7 +222,7 @@ let compute_new_princ_type_from_rel rel_to_fun sorts princ_type =
in
let pre_res =
replace_vars
- (list_map_i (fun i id -> (id, mkRel i)) 1 ptes_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
@@ -271,8 +238,10 @@ 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)
+ let args,ty = decompose_prod t in
+ let s = destSort ty in
+ Global.add_constraints (Univ.enforce_leq (univ_of_sort toSort) (univ_of_sort s) Univ.Constraint.empty);
+ compose_prod args (mkSort toSort)
)
in
let princName_as_constr = Constrintern.global_reference princName in
@@ -288,23 +257,6 @@ let change_property_sort toSort princ princName =
)
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 e when Errors.noncritical e -> mt ()
- ) ++msg)
-
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
@@ -319,23 +271,25 @@ let build_functional_principle interactive_proof old_princ_type sorts funs i pro
(* 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_________") []
+ next_ident_away_in_goal (Id.of_string "___________princ_________") []
in
+ let hook = Lemmas.mk_hook (hook new_principle_type) in
begin
Lemmas.start_proof
new_princ_name
- (Decl_kinds.Global,(Decl_kinds.Proof Decl_kinds.Theorem))
- new_principle_type
- (hook new_principle_type)
+ (Decl_kinds.Global,false,(Decl_kinds.Proof Decl_kinds.Theorem))
+ (*FIXME*) Evd.empty
+ new_principle_type
+ hook
;
(* let _tim1 = System.get_time () in *)
- Pfedit.by (proof_tac (Array.map mkConst funs) mutr_nparams);
+ ignore (Pfedit.by (Proofview.V82.tactic (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
+ get_proof_clean true, Ephemeron.create hook
end
@@ -347,7 +301,7 @@ let generate_functional_principle
try
let f = funs.(i) in
- let type_sort = Termops.new_sort_in_family InType in
+ let type_sort = Universes.new_sort_in_family InType in
let new_sorts =
match sorts with
| None -> Array.make (Array.length funs) (type_sort)
@@ -357,42 +311,35 @@ let generate_functional_principle
match new_princ_name with
| Some (id) -> id,id
| None ->
- let id_of_f = id_of_label (con_label f) in
+ let id_of_f = Label.to_id (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
+ let hook new_principle_type _ _ =
+ if Option.is_empty sorts
then
- (* let id_of_f = id_of_label (con_label f) in *)
+ (* let id_of_f = Label.to_id (con_label f) in *)
let register_with_sort fam_sort =
- let s = Termops.new_sort_in_family fam_sort in
+ let s = Universes.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_secctx = None;
- const_entry_type = None;
- const_entry_opaque = false }
- in
+ let ce = Declare.definition_entry value in (*FIXME, no poly, nothing *)
ignore(
Declare.declare_constant
name
(Entries.DefinitionEntry ce,
- Decl_kinds.IsDefinition (Decl_kinds.Scheme)
- )
+ Decl_kinds.IsDefinition (Decl_kinds.Scheme))
);
- Flags.if_verbose
- (fun id -> Pp.msgnl (Ppconstr.pr_id id ++ str " is defined"))
- name;
+ Declare.definition_message 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
+ 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 !!!!
@@ -403,10 +350,10 @@ let generate_functional_principle
begin
try
let id = Pfedit.get_current_proof_name () in
- let s = string_of_id id in
+ let s = Id.to_string id in
let n = String.length "___________princ_________" in
if String.length s >= n
- then if String.sub s 0 n = "___________princ_________"
+ then if String.equal (String.sub s 0 n) "___________princ_________"
then Pfedit.delete_current_proof ()
else ()
else ()
@@ -420,26 +367,25 @@ let generate_functional_principle
exception Not_Rec
let get_funs_constant mp dp =
- let rec get_funs_constant const e : (Names.constant*int) array =
+ let 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
+ let const = make_con mp dp (Label.of_id id) in
const,i
| Anonymous ->
- anomaly "Anonymous fix"
+ anomaly (Pp.str "Anonymous fix")
)
na
| _ -> [|const,0|]
in
function const ->
let find_constant_body const =
- match body_of_constant (Global.lookup_constant const) with
- | Some b ->
- let body = force b in
+ match Global.body_of_constant const with
+ | Some body ->
let body = Tacred.cbv_norm_flags
(Closure.RedFlags.mkflags [Closure.RedFlags.fZETA])
(Global.env ())
@@ -462,7 +408,7 @@ let get_funs_constant mp dp =
let first_params = List.hd l_params in
List.iter
(fun params ->
- if not (list_equal (fun (n1, c1) (n2, c2) -> n1 = n2 && eq_constr c1 c2) first_params params)
+ if not (List.equal (fun (n1, c1) (n2, c2) -> Name.equal n1 n2 && eq_constr c1 c2) first_params params)
then error "Not a mutal recursive block"
)
l_params
@@ -474,14 +420,15 @@ let get_funs_constant mp dp =
match kind_of_term body with
| Fix((idxs,_),(na,ta,ca)) -> (idxs,na,ta,ca)
| _ ->
- if is_first && (List.length l_bodies = 1)
+ if is_first && Int.equal (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 *)
let eq_infos (ia1, na1, ta1, ca1) (ia2, na2, ta2, ca2) =
- ia1 = ia2 && na1 = na2 && array_equal eq_constr ta1 ta2 && array_equal eq_constr ca1 ca2
+ Array.equal Int.equal ia1 ia2 && Array.equal Name.equal na1 na2 &&
+ Array.equal eq_constr ta1 ta2 && Array.equal eq_constr ca1 ca2
in
if not (eq_infos first_infos (extract_info false body))
then error "Not a mutal recursive block"
@@ -494,7 +441,7 @@ let get_funs_constant mp dp =
exception No_graph_found
exception Found_type of int
-let make_scheme (fas : (constant*Glob_term.glob_sort) list) : Entries.definition_entry list =
+let make_scheme (fas : (constant*glob_sort) list) : Entries.definition_entry list =
let env = Global.env ()
and sigma = Evd.empty in
let funs = List.map fst fas in
@@ -513,26 +460,27 @@ let make_scheme (fas : (constant*Glob_term.glob_sort) list) : Entries.definition
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)
+ (function cst -> List.assoc_f Constant.equal cst this_block_funs_indexes)
funs
in
let ind_list =
List.map
(fun (idx) ->
let ind = first_fun_kn,idx in
- ind,true,prop_sort
+ (ind,Univ.Instance.empty)(*FIXME*),true,prop_sort
)
funs_indexes
in
+ let sigma, schemes =
+ Indrec.build_mutual_induction_scheme env sigma ind_list
+ in
let l_schemes =
- List.map
- (Typing.type_of env sigma)
- (Indrec.build_mutual_induction_scheme env sigma ind_list)
+ List.map (Typing.type_of env sigma) schemes
in
let i = ref (-1) in
let sorts =
List.rev_map (fun (_,x) ->
- Termops.new_sort_in_family (Pretyping.interp_elimination_sort x)
+ Universes.new_sort_in_family (Pretyping.interp_elimination_sort x)
)
fas
in
@@ -540,9 +488,9 @@ let make_scheme (fas : (constant*Glob_term.glob_sort) list) : Entries.definition
let first_type,other_princ_types =
match l_schemes with
s::l_schemes -> s,l_schemes
- | _ -> anomaly ""
+ | _ -> anomaly (Pp.str "")
in
- let (_,(const,_,_)) =
+ let ((_,(const,_)),_) =
try
build_functional_principle false
first_type
@@ -556,10 +504,10 @@ let make_scheme (fas : (constant*Glob_term.glob_sort) list) : Entries.definition
begin
try
let id = Pfedit.get_current_proof_name () in
- let s = string_of_id id in
+ let s = Id.to_string id in
let n = String.length "___________princ_________" in
if String.length s >= n
- then if String.sub s 0 n = "___________princ_________"
+ then if String.equal (String.sub s 0 n) "___________princ_________"
then Pfedit.delete_current_proof ()
else ()
else ()
@@ -574,13 +522,13 @@ let make_scheme (fas : (constant*Glob_term.glob_sort) list) : Entries.definition
let finfos = find_Function_infos this_block_funs.(0) in
try
let equation = Option.get finfos.equation_lemma in
- Declarations.is_opaque (Global.lookup_constant equation)
+ Declareops.is_opaque (Global.lookup_constant equation)
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 = []
+ if List.is_empty other_princ_types
then
[const]
else
@@ -590,7 +538,7 @@ let make_scheme (fas : (constant*Glob_term.glob_sort) list) : Entries.definition
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 ctxt,fix = decompose_lam_assum (fst(fst(Future.force 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 *)
@@ -616,7 +564,7 @@ let make_scheme (fas : (constant*Glob_term.glob_sort) list) : Entries.definition
(* If we reach this point, the two principle are not mutually recursive
We fall back to the previous method
*)
- let (_,(const,_,_)) =
+ let ((_,(const,_)),_) =
build_functional_principle
false
(List.nth other_princ_types (!i - 1))
@@ -632,7 +580,8 @@ let make_scheme (fas : (constant*Glob_term.glob_sort) list) : Entries.definition
Termops.it_mkLambda_or_LetIn (mkFix((idxs,i),decl)) ctxt
in
{const with
- Entries.const_entry_body = princ_body;
+ Entries.const_entry_body =
+ (Future.from_val (Term_typing.mk_pure_proof princ_body));
Entries.const_entry_type = Some scheme_type
}
)
@@ -648,11 +597,11 @@ let build_scheme fas =
(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"
+ match Smartlocate.global_with_alias f with
+ | Globnames.ConstRef c -> c
+ | _ -> Errors.error "Functional Scheme can only be used with functions"
with Not_found ->
- Util.error ("Cannot find "^ Libnames.string_of_reference f)
+ Errors.error ("Cannot find "^ Libnames.string_of_reference f)
in
(f_as_constant,sort)
)
@@ -665,8 +614,7 @@ let build_scheme fas =
(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
+ Declare.definition_message princ_id
)
fas
bodies_types;
@@ -681,10 +629,10 @@ let build_case_scheme fa =
(* Constrintern.global_reference id *)
(* in *)
let funs = (fun (_,f,_) ->
- try Libnames.constr_of_global (Nametab.global f)
+ try fst (Universes.unsafe_constr_of_global (Smartlocate.global_with_alias f))
with Not_found ->
- Util.error ("Cannot find "^ Libnames.string_of_reference f)) fa in
- let first_fun = destConst funs in
+ Errors.error ("Cannot find "^ Libnames.string_of_reference f)) fa in
+ let first_fun,u = 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
@@ -696,16 +644,18 @@ let build_case_scheme fa =
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
+ List.assoc_f Constant.equal (fst (destConst funs)) this_block_funs_indexes
in
let ind_fun =
let ind = first_fun_kn,funs_indexes in
- ind,prop_sort
+ (ind,Univ.Instance.empty)(*FIXME*),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 sigma, scheme =
+ (fun (ind,sf) -> Indrec.build_case_analysis_scheme_default env sigma ind sf) ind_fun in
+ let scheme_type = (Typing.type_of env sigma ) scheme in
let sorts =
(fun (_,_,x) ->
- Termops.new_sort_in_family (Pretyping.interp_elimination_sort x)
+ Universes.new_sort_in_family (Pretyping.interp_elimination_sort x)
)
fa
in
@@ -722,6 +672,6 @@ let build_case_scheme fa =
(Some princ_name)
this_block_funs
0
- (prove_princ_for_struct false 0 [|destConst funs|])
+ (prove_princ_for_struct false 0 [|fst (destConst funs)|])
in
()
diff --git a/plugins/funind/functional_principles_types.mli b/plugins/funind/functional_principles_types.mli
index 1c02c16e..a16b834f 100644
--- a/plugins/funind/functional_principles_types.mli
+++ b/plugins/funind/functional_principles_types.mli
@@ -1,6 +1,6 @@
open Names
open Term
-
+open Misctypes
val generate_functional_principle :
(* do we accept interactive proving *)
@@ -10,7 +10,7 @@ val generate_functional_principle :
(* *)
sorts array option ->
(* Name of the new principle *)
- (identifier) option ->
+ (Id.t) option ->
(* the compute functions to use *)
constant array ->
(* We prove the nth- principle *)
@@ -27,8 +27,8 @@ val compute_new_princ_type_from_rel : constr array -> sorts array ->
exception No_graph_found
-val make_scheme : (constant*Glob_term.glob_sort) list -> Entries.definition_entry list
+val make_scheme : (constant*glob_sort) list -> Entries.definition_entry list
-val build_scheme : (identifier*Libnames.reference*Glob_term.glob_sort) list -> unit
-val build_case_scheme : (identifier*Libnames.reference*Glob_term.glob_sort) -> unit
+val build_scheme : (Id.t*Libnames.reference*glob_sort) list -> unit
+val build_case_scheme : (Id.t*Libnames.reference*glob_sort) -> unit
diff --git a/plugins/funind/g_indfun.ml4 b/plugins/funind/g_indfun.ml4
index ffaa2208..fd48ab59 100644
--- a/plugins/funind/g_indfun.ml4
+++ b/plugins/funind/g_indfun.ml4
@@ -1,35 +1,38 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
+(*i camlp4deps: "grammar/grammar.cma" i*)
+open Compat
open Util
open Term
+open Vars
open Names
open Pp
-open Topconstr
+open Constrexpr
open Indfun_common
open Indfun
open Genarg
-open Pcoq
open Tacticals
-open Constr
+open Misctypes
+
+DECLARE PLUGIN "recdef_plugin"
let pr_binding prc = function
- | loc, Glob_term.NamedHyp id, c -> hov 1 (Ppconstr.pr_id id ++ str " := " ++ cut () ++ prc c)
- | loc, Glob_term.AnonHyp n, c -> hov 1 (int n ++ str " := " ++ cut () ++ prc c)
+ | loc, NamedHyp id, c -> hov 1 (Ppconstr.pr_id id ++ str " := " ++ cut () ++ prc c)
+ | loc, AnonHyp n, c -> hov 1 (int n ++ str " := " ++ cut () ++ prc c)
let pr_bindings prc prlc = function
- | Glob_term.ImplicitBindings l ->
+ | ImplicitBindings l ->
brk (1,1) ++ str "with" ++ brk (1,1) ++
- Util.prlist_with_sep spc prc l
- | Glob_term.ExplicitBindings l ->
+ pr_sequence prc l
+ | ExplicitBindings l ->
brk (1,1) ++ str "with" ++ brk (1,1) ++
- Util.prlist_with_sep spc (fun b -> str"(" ++ pr_binding prlc b ++ str")") l
- | Glob_term.NoBindings -> mt ()
+ 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)
@@ -69,18 +72,23 @@ END
TACTIC EXTEND newfuninv
[ "functional" "inversion" quantified_hypothesis(hyp) reference_opt(fname) ] ->
[
- Invfun.invfun hyp fname
+ Proofview.V82.tactic (Invfun.invfun hyp fname)
]
END
-let pr_intro_as_pat prc _ _ pat =
+let pr_intro_as_pat _prc _ _ pat =
match pat with
- | Some pat -> spc () ++ str "as" ++ spc () ++ pr_intro_pattern pat
+ | Some pat ->
+ spc () ++ str "as" ++ spc () ++ (* Miscprint.pr_intro_pattern prc pat *)
+ str"<simple_intropattern>"
| None -> mt ()
+let out_disjunctive = function
+ | loc, IntroAction (IntroOrAndPattern l) -> (loc,l)
+ | _ -> Errors.error "Disjunctive or conjunctive intro pattern expected."
-ARGUMENT EXTEND with_names TYPED AS intro_pattern_opt PRINTED BY pr_intro_as_pat
+ARGUMENT EXTEND with_names TYPED AS simple_intropattern_opt PRINTED BY pr_intro_as_pat
| [ "as" simple_intropattern(ipat) ] -> [ Some ipat ]
| [] ->[ None ]
END
@@ -96,7 +104,7 @@ TACTIC EXTEND newfunind
| [c] -> c
| c::cl -> applist(c,cl)
in
- Extratactics.onSomeWithHoles (fun x -> functional_induction true c x pat) princl ]
+ Extratactics.onSomeWithHoles (fun x -> Proofview.V82.tactic (functional_induction true c x (Option.map out_disjunctive pat))) princl ]
END
(***** debug only ***)
TACTIC EXTEND snewfunind
@@ -107,11 +115,11 @@ TACTIC EXTEND snewfunind
| [c] -> c
| c::cl -> applist(c,cl)
in
- Extratactics.onSomeWithHoles (fun x -> functional_induction false c x pat) princl ]
+ Extratactics.onSomeWithHoles (fun x -> Proofview.V82.tactic (functional_induction false c x (Option.map out_disjunctive pat))) princl ]
END
-let pr_constr_coma_sequence prc _ _ = Util.prlist_with_sep Util.pr_comma prc
+let pr_constr_coma_sequence prc _ _ = prlist_with_sep pr_comma prc
ARGUMENT EXTEND constr_coma_sequence'
TYPED AS constr_list
@@ -133,34 +141,37 @@ module Gram = Pcoq.Gram
module Vernac = Pcoq.Vernac_
module Tactic = Pcoq.Tactic
-module FunctionGram =
-struct
- let gec s = Gram.entry_create ("Function."^s)
- (* types *)
- let function_rec_definition_loc : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) located Gram.entry = gec "function_rec_definition_loc"
-end
-open FunctionGram
+type function_rec_definition_loc_argtype = (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) Loc.located
+
+let (wit_function_rec_definition_loc : function_rec_definition_loc_argtype Genarg.uniform_genarg_type) =
+ Genarg.create_arg None "function_rec_definition_loc"
+
+let function_rec_definition_loc =
+ Pcoq.create_generic_entry "function_rec_definition_loc" (Genarg.rawwit wit_function_rec_definition_loc)
GEXTEND Gram
GLOBAL: function_rec_definition_loc ;
function_rec_definition_loc:
- [ [ g = Vernac.rec_definition -> loc, g ]]
+ [ [ g = Vernac.rec_definition -> !@loc, g ]]
;
- END
-type 'a function_rec_definition_loc_argtype = ((Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) located, 'a) Genarg.abstract_argument_type
+END
-let (wit_function_rec_definition_loc : Genarg.tlevel function_rec_definition_loc_argtype),
- (globwit_function_rec_definition_loc : Genarg.glevel function_rec_definition_loc_argtype),
- (rawwit_function_rec_definition_loc : Genarg.rlevel function_rec_definition_loc_argtype) =
- Genarg.create_arg None "function_rec_definition_loc"
+(* TASSI: n'importe quoi ! *)
VERNAC COMMAND EXTEND Function
- ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")] ->
- [
- do_generate_principle false (List.map snd recsl);
-
- ]
+ ["Function" ne_function_rec_definition_loc_list_sep(recsl,"with")]
+ => [ let hard = List.exists (function
+ | _,((_,(_,(CMeasureRec _|CWfRec _)),_,_,_),_) -> true
+ | _,((_,(_,CStructRec),_,_,_),_) -> false) recsl in
+ match
+ Vernac_classifier.classify_vernac
+ (Vernacexpr.VernacFixpoint(None, List.map snd recsl))
+ with
+ | Vernacexpr.VtSideff ids, _ when hard ->
+ Vernacexpr.(VtStartProof ("Classic", GuaranteesOpacity, ids), VtLater)
+ | x -> x ]
+ -> [ do_generate_principle false (List.map snd recsl) ]
END
let pr_fun_scheme_arg (princ_name,fun_name,s) =
@@ -175,23 +186,25 @@ END
let warning_error names e =
- let e = Cerrors.process_vernac_interp_error e in
+ let (e, _) = Cerrors.process_vernac_interp_error (e, Exninfo.null) in
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) ++
+ h 1 (pr_enum Libnames.pr_reference names) ++
if do_observe () then (spc () ++ Errors.print 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) ++
+ h 1 (pr_enum Libnames.pr_reference names) ++
if do_observe () then Errors.print e else mt ())
| _ -> raise e
VERNAC COMMAND EXTEND NewFunctionalScheme
- ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ] ->
+ ["Functional" "Scheme" ne_fun_scheme_arg_list_sep(fas,"with") ]
+ => [ Vernacexpr.VtSideff(List.map pi1 fas), Vernacexpr.VtLater ]
+ ->
[
begin
try
@@ -202,13 +215,13 @@ VERNAC COMMAND EXTEND NewFunctionalScheme
| (_,fun_name,_)::_ ->
begin
begin
- make_graph (Nametab.global fun_name)
+ make_graph (Smartlocate.global_with_alias 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 when Errors.noncritical e ->
+ Errors.error ("Cannot generate induction principle(s)")
+ | e when Errors.noncritical e ->
let names = List.map (fun (_,na,_) -> na) fas in
warning_error names e
@@ -225,15 +238,14 @@ END
(***** debug only ***)
VERNAC COMMAND EXTEND NewFunctionalCase
- ["Functional" "Case" fun_scheme_arg(fas) ] ->
- [
- Functional_principles_types.build_case_scheme fas
- ]
+ ["Functional" "Case" fun_scheme_arg(fas) ]
+ => [ Vernacexpr.VtSideff[pi1 fas], Vernacexpr.VtLater ]
+ -> [ Functional_principles_types.build_case_scheme fas ]
END
(***** debug only ***)
-VERNAC COMMAND EXTEND GenerateGraph
-["Generate" "graph" "for" reference(c)] -> [ make_graph (Nametab.global c) ]
+VERNAC COMMAND EXTEND GenerateGraph CLASSIFIED AS QUERY
+["Generate" "graph" "for" reference(c)] -> [ make_graph (Smartlocate.global_with_alias c) ]
END
@@ -273,7 +285,7 @@ let constr_head_match u t=
if isApp u
then
let uhd,args= destApp u in
- uhd=t
+ Constr.equal uhd t
else false
(** [hdMatchSub inu t] returns the list of occurrences of [t] in
@@ -296,22 +308,25 @@ let rec hdMatchSub inu (test: constr -> bool) : fapp_info list =
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;
+ let max_rel = try Int.Set.max_elt freeset with Not_found -> -1 in
+ {fname = f; largs = args; free = Int.Set.is_empty freeset;
max_rel = max_rel; onlyvars = List.for_all isVar args }
::subres
+let make_eq () =
+(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq ())
+
let mkEq typ c1 c2 =
- mkApp (Coqlib.build_coq_eq(),[| typ; c1; c2|])
+ mkApp (make_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)
+ (Proofview.V82.of_tactic (Tactics.letin_tac None (Name idunsafe) cstr None Locusops.allHypsAndConcl))
(tclTHENFIRST
- (Tactics.assert_tac Anonymous (mkEq typ (mkVar idunsafe) cstr))
- Tactics.reflexivity)
+ (Proofview.V82.of_tactic (Tactics.assert_before Anonymous (mkEq typ (mkVar idunsafe) cstr)))
+ (Proofview.V82.of_tactic Tactics.reflexivity))
gl
@@ -357,7 +372,7 @@ let poseq_list_ids lcstr gl =
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
+ List.fold_right (List.add_set Pervasives.(=)) pre_res [] in
(prlistconstr (List.map (fun x -> applist (x.fname,x.largs)) res);
res)
@@ -367,7 +382,7 @@ let find_fapp (test:constr -> bool) g : fapp_info list =
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)
+let finduction (oid:Id.t option) (heuristic: fapp_info list -> fapp_info list)
(nexttac:Proof_type.tactic) g =
let test = match oid with
| Some id ->
@@ -377,7 +392,7 @@ let finduction (oid:identifier option) (heuristic: fapp_info list -> fapp_info l
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";
+ if List.is_empty ordered_info_list then Errors.error "function not found in goal\n";
let taclist: Proof_type.tactic list =
List.map
(fun info ->
@@ -419,10 +434,10 @@ TACTIC EXTEND finduction
["finduction" ident(id) natural_opt(oi)] ->
[
match oi with
- | Some(n) when n<=0 -> Util.error "numerical argument must be > 0"
+ | Some(n) when n<=0 -> Errors.error "numerical argument must be > 0"
| _ ->
let heuristic = chose_heuristic oi in
- finduction (Some id) heuristic tclIDTAC
+ Proofview.V82.tactic (finduction (Some id) heuristic tclIDTAC)
]
END
@@ -432,13 +447,13 @@ TACTIC EXTEND fauto
[ "fauto" tactic(tac)] ->
[
let heuristic = chose_heuristic None in
- finduction None heuristic (Tacinterp.eval_tactic tac)
+ Proofview.V82.tactic (finduction None heuristic (Proofview.V82.of_tactic (Tacinterp.eval_tactic tac)))
]
|
[ "fauto" ] ->
[
let heuristic = chose_heuristic None in
- finduction None heuristic tclIDTAC
+ Proofview.V82.tactic (finduction None heuristic tclIDTAC)
]
END
@@ -446,31 +461,31 @@ END
TACTIC EXTEND poseq
[ "poseq" ident(x) constr(c) ] ->
- [ poseq x c ]
+ [ Proofview.V82.tactic (poseq x c) ]
END
-VERNAC COMMAND EXTEND Showindinfo
+VERNAC COMMAND EXTEND Showindinfo CLASSIFIED AS QUERY
[ "showindinfo" ident(x) ] -> [ Merge.showind x ]
END
-VERNAC COMMAND EXTEND MergeFunind
+VERNAC COMMAND EXTEND MergeFunind CLASSIFIED AS SIDEFF
[ "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 f1,ctx = Constrintern.interp_constr (Global.env()) Evd.empty
+ (CRef (Libnames.Ident (Loc.ghost,id1),None)) in
+ let f2,ctx' = Constrintern.interp_constr (Global.env()) Evd.empty
+ (CRef (Libnames.Ident (Loc.ghost,id2),None)) 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
+ if not (Int.equal ar1 (List.length cl1)) then
+ Errors.error ("not the right number of arguments for " ^ Id.to_string id1) in
let _ =
- if ar2 <> List.length cl2 then
- Util.error ("not the right number of arguments for " ^ string_of_id id2) in
+ if not (Int.equal ar2 (List.length cl2)) then
+ Errors.error ("not the right number of arguments for " ^ Id.to_string id2) in
Merge.merge id1 id2 (Array.of_list cl1) (Array.of_list cl2) id
]
END
diff --git a/plugins/funind/glob_term_to_relation.ml b/plugins/funind/glob_term_to_relation.ml
index b9e0e62a..a2577e2b 100644
--- a/plugins/funind/glob_term_to_relation.ml
+++ b/plugins/funind/glob_term_to_relation.ml
@@ -2,26 +2,30 @@ open Printer
open Pp
open Names
open Term
+open Vars
open Glob_term
-open Libnames
+open Glob_ops
+open Globnames
open Indfun_common
+open Errors
open Util
open Glob_termops
+open Misctypes
let observe strm =
if do_observe ()
- then Pp.msgnl strm
+ then Pp.msg_debug strm
else ()
-let observennl strm =
+(*let observennl strm =
if do_observe ()
then Pp.msg strm
- else ()
+ else ()*)
type binder_type =
- | Lambda of name
- | Prod of name
- | LetIn of name
+ | Lambda of Name.t
+ | Prod of Name.t
+ | LetIn of Name.t
type glob_context = (binder_type*glob_constr) list
@@ -54,7 +58,7 @@ type 'a build_entry_pre_return =
type 'a build_entry_return =
{
result : 'a build_entry_pre_return list;
- to_avoid : identifier list
+ to_avoid : Id.t list
}
(*
@@ -86,7 +90,7 @@ let combine_results
in (* and then we flatten the map *)
{
result = List.concat pre_result;
- to_avoid = list_union res1.to_avoid res2.to_avoid
+ to_avoid = List.union Id.equal res1.to_avoid res2.to_avoid
}
@@ -111,9 +115,9 @@ let ids_of_binder = function
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
+ let new_mapping = List.fold_right Id.Map.remove (ids_of_binder bt) mapping in
(bt,change_vars mapping t)::
- (if idmap_is_empty new_mapping
+ (if Id.Map.is_empty new_mapping
then l
else change_vars_in_binder new_mapping l
)
@@ -122,7 +126,7 @@ 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)
+ if Id.List.mem x_id (ids_of_binder bt)
then l
else replace_var_by_term_in_binder x_id term l
@@ -130,28 +134,28 @@ 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
+ List.exists (is_free_in id) args || Id.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) =
+ let next_name_away (na:Name.t) (mapping: Id.t Id.Map.t) (avoid: Id.t list) =
match na with
- | Name id when List.mem id avoid ->
+ | Name id when Id.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
+ Name new_id,Id.Map.add id new_id mapping,new_id::avoid
| _ -> na,mapping,avoid
in
- let next_bt_away bt (avoid:identifier list) =
+ let next_bt_away bt (avoid:Id.t list) =
match bt with
| LetIn na ->
- let new_na,mapping,new_avoid = next_name_away na Idmap.empty avoid in
+ let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in
LetIn new_na,mapping,new_avoid
| Prod na ->
- let new_na,mapping,new_avoid = next_name_away na Idmap.empty avoid in
+ let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in
Prod new_na,mapping,new_avoid
| Lambda na ->
- let new_na,mapping,new_avoid = next_name_away na Idmap.empty avoid in
+ let new_na,mapping,new_avoid = next_name_away na Id.Map.empty avoid in
Lambda new_na,mapping,new_avoid
in
let rec do_apply avoid ctxt body args =
@@ -170,7 +174,7 @@ let apply_args ctxt body args =
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 mapping = Id.Map.add id new_id Id.Map.empty in
let new_ctxt' = change_vars_in_binder mapping ctxt' in
let new_body = change_vars mapping body in
new_avoid',new_ctxt',new_body,new_id
@@ -266,11 +270,11 @@ let make_discr_match_el =
end
*)
let make_discr_match_brl i =
- list_map_i
+ List.map_i
(fun j (_,idl,patl,_) ->
- if j=i
- then (dummy_loc,idl,patl, mkGRef (Lazy.force coq_True_ref))
- else (dummy_loc,idl,patl, mkGRef (Lazy.force coq_False_ref))
+ if Int.equal j i
+ then (Loc.ghost,idl,patl, mkGRef (Lazy.force coq_True_ref))
+ else (Loc.ghost,idl,patl, mkGRef (Lazy.force coq_False_ref))
)
0
(*
@@ -285,10 +289,6 @@ let make_discr_match brl =
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 *)
(**********************************************************************)
@@ -304,18 +304,17 @@ let build_constructors_of_type ind' argl =
Impargs.implicits_of_global constructref
in
let cst_narg =
- Inductiveops.mis_constructor_nargs_env
+ Inductiveops.constructor_nallargs_env
(Global.env ())
construct
in
- let argl = match argl with
- | None ->
+ let argl =
+ if List.is_empty argl
+ then
Array.to_list
- (Array.init cst_narg (fun _ -> mkGHole ())
+ (Array.init (cst_narg - npar) (fun _ -> mkGHole ())
)
- | Some l ->
- Array.to_list
- (Array.init npar (fun _ -> mkGHole ()))@l
+ else argl
in
let pat_as_term =
mkGApp(mkGRef (ConstructRef(ind',i+1)),argl)
@@ -324,40 +323,6 @@ let build_constructors_of_type ind' argl =
)
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,_ = glob_decompose_app b in
- match f with
- | GRef(_,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
- | GCast(_,b,_) -> find_type_of nb b
- | GApp _ -> assert false (* we have decomposed any application via glob_decompose_app *)
- | _ -> raise (Invalid_argument "not a ref")
-
-
-
-
(******************)
(* Main functions *)
(******************)
@@ -368,14 +333,14 @@ 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
+ let value = Option.map (fun x-> fst (Pretyping.understand env Evd.empty x)) raw_value in
+ let typ,ctx = Pretyping.understand env Evd.empty ~expected_type:Pretyping.IsType 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);
+ observe (str "new rel env := " ++ Printer.pr_rel_context_of env Evd.empty);
match pat with
| PatVar(_,na) -> Environ.push_rel (na,None,typ) env
@@ -385,14 +350,14 @@ let add_pat_variables pat typ env : Environ.env =
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 constructor : Inductiveops.constructor_summary = List.find (fun cs -> eq_constructor c (fst cs.Inductiveops.cs_cstr)) (Array.to_list constructors) in
let cs_args_types :types list = List.map (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
+ Context.fold_rel_context
(fun (na,v,t) (env,ctxt) ->
match na with
| Anonymous -> assert false
@@ -411,7 +376,7 @@ let add_pat_variables pat typ env : Environ.env =
~init:(env,[])
)
in
- observe (str "new var env := " ++ Printer.pr_named_context_of res);
+ observe (str "new var env := " ++ Printer.pr_named_context_of res Evd.empty);
res
@@ -423,7 +388,7 @@ let rec pattern_to_term_and_type env typ = function
mkGVar id
| PatCstr(loc,constr,patternl,_) ->
let cst_narg =
- Inductiveops.mis_constructor_nargs_env
+ Inductiveops.constructor_nallargs_env
(Global.env ())
constr
in
@@ -432,7 +397,7 @@ let rec pattern_to_term_and_type env typ = function
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 constructor = List.find (fun cs -> eq_constructor (fst 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
@@ -440,7 +405,7 @@ let rec pattern_to_term_and_type env typ = function
Array.to_list
(Array.init
(cst_narg - List.length patternl)
- (fun i -> Detyping.detype false [] (Termops.names_of_rel_context env) csta.(i))
+ (fun i -> Detyping.detype false [] env Evd.empty csta.(i))
)
in
let patl_as_term =
@@ -508,12 +473,12 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
| u::l ->
match t with
| GLambda(loc,na,_,nat,b) ->
- GLetIn(dummy_loc,na,u,aux b l)
+ GLetIn(Loc.ghost,na,u,aux b l)
| _ ->
- GApp(dummy_loc,t,l)
+ GApp(Loc.ghost,t,l)
in
build_entry_lc env funnames avoid (aux f args)
- | GVar(_,id) when Idset.mem id funnames ->
+ | GVar(_,id) when Id.Set.mem id funnames ->
(* if we have [f t1 ... tn] with [f]$\in$[fnames]
then we create a fresh variable [res],
add [res] and its "value" (i.e. [res v1 ... vn]) to each
@@ -521,10 +486,10 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
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_as_constr,ctx = Pretyping.understand env Evd.empty 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 res_raw_type = Detyping.detype false [] env Evd.empty rt_typ in
+ let res = fresh_id args_res.to_avoid "_res" in
let new_avoid = res::args_res.to_avoid in
let res_rt = mkGVar res in
let new_result =
@@ -568,7 +533,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let new_b =
replace_var_by_term
id
- (GVar(dummy_loc,id))
+ (GVar(Loc.ghost,id))
b
in
(Name new_id,new_b,new_avoid)
@@ -629,7 +594,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
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_as_constr,ctx = Pretyping.understand env Evd.empty v in
let v_type = Typing.type_of env Evd.empty v_as_constr in
let new_env =
match n with
@@ -645,7 +610,7 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
let make_discr = make_discr_match brl in
build_entry_lc_from_case env funnames make_discr el brl avoid
| GIf(_,b,(na,e_option),lhs,rhs) ->
- let b_as_constr = Pretyping.Default.understand Evd.empty env b in
+ let b_as_constr,ctx = Pretyping.understand env Evd.empty 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
@@ -654,11 +619,11 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
Printer.pr_glob_constr b ++ str " in " ++
Printer.pr_glob_constr rt ++ str ". try again with a cast")
in
- let case_pats = build_constructors_of_type ind None in
- assert (Array.length case_pats = 2);
+ let case_pats = build_constructors_of_type (fst ind) [] in
+ assert (Int.equal (Array.length case_pats) 2);
let brl =
- list_map_i
- (fun i x -> (dummy_loc,[],[case_pats.(i)],x))
+ List.map_i
+ (fun i x -> (Loc.ghost,[],[case_pats.(i)],x))
0
[lhs;rhs]
in
@@ -670,14 +635,14 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
| GLetTuple(_,nal,_,b,e) ->
begin
let nal_as_glob_constr =
- Some (List.map
+ List.map
(function
Name id -> mkGVar id
| Anonymous -> mkGHole ()
)
- nal)
+ nal
in
- let b_as_constr = Pretyping.Default.understand Evd.empty env b in
+ let b_as_constr,ctx = Pretyping.understand env Evd.empty 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
@@ -686,10 +651,10 @@ let rec build_entry_lc env funnames avoid rt : glob_constr build_entry_return =
Printer.pr_glob_constr b ++ str " in " ++
Printer.pr_glob_constr rt ++ str ". try again with a cast")
in
- let case_pats = build_constructors_of_type ind nal_as_glob_constr in
- assert (Array.length case_pats = 1);
+ let case_pats = build_constructors_of_type (fst ind) nal_as_glob_constr in
+ assert (Int.equal (Array.length case_pats) 1);
let br =
- (dummy_loc,[],[case_pats.(0)],e)
+ (Loc.ghost,[],[case_pats.(0)],e)
in
let match_expr = mkGCases(None,[b,(Anonymous,None)],[br]) in
build_entry_lc env funnames avoid match_expr
@@ -724,7 +689,7 @@ and build_entry_lc_from_case env funname make_discr
in
let types =
List.map (fun (case_arg,_) ->
- let case_arg_as_constr = Pretyping.Default.understand Evd.empty env case_arg in
+ let case_arg_as_constr,ctx = Pretyping.understand env Evd.empty case_arg in
Typing.type_of env Evd.empty case_arg_as_constr
) el
in
@@ -746,7 +711,8 @@ and build_entry_lc_from_case env funname make_discr
{
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
+ List.fold_left (fun acc r -> List.union Id.equal acc r.to_avoid)
+ [] results
}
and build_entry_lc_from_case_term env types funname make_discr patterns_to_prevent brl avoid
@@ -761,7 +727,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
(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 -> glob_constr -> glob_constr) list =
+ let not_those_patterns : (Id.t list -> glob_constr -> glob_constr) list =
List.map2
(fun pat typ ->
fun avoid pat'_as_term ->
@@ -775,7 +741,7 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
in
let raw_typ_of_id =
Detyping.detype false []
- (Termops.names_of_rel_context env_with_pat_ids) typ_of_id
+ env_with_pat_ids Evd.empty typ_of_id
in
mkGProd (Name id,raw_typ_of_id,acc))
pat_ids
@@ -816,18 +782,18 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
let those_pattern_preconds =
(List.flatten
(
- list_map3
+ 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 typ = Detyping.detype false [] new_env Evd.empty 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
+ if Id.Set.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
+ Detyping.detype false [] new_env Evd.empty typ_of_id
in
raw_typ_of_id
)::acc
@@ -871,14 +837,14 @@ and build_entry_lc_from_case_term env types funname make_discr patterns_to_preve
let is_res id =
try
- String.sub (string_of_id id) 0 3 = "res"
+ String.equal (String.sub (Id.to_string id) 0 4) "_res"
with Invalid_argument _ -> false
let same_raw_term rt1 rt2 =
match rt1,rt2 with
- | GRef(_,r1), GRef (_,r2) -> r1=r2
+ | GRef(_,r1,_), GRef (_,r2,_) -> Globnames.eq_gr r1 r2
| GHole _, GHole _ -> true
| _ -> false
let decompose_raw_eq lhs rhs =
@@ -892,7 +858,7 @@ let decompose_raw_eq lhs rhs =
observe (str "lrhs := " ++ int (List.length lrhs));
let sllhs = List.length llhs in
let slrhs = List.length lrhs in
- if same_raw_term lhd rhd && sllhs = slrhs
+ if same_raw_term lhd rhd && Int.equal sllhs slrhs
then
(* let _ = assert false in *)
List.fold_right2 decompose_raw_eq llhs lrhs acc
@@ -928,7 +894,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
let new_t =
mkGApp(mkGVar(mk_rel_id this_relname),args'@[res_rt])
in
- let t' = Pretyping.Default.understand Evd.empty env new_t in
+ let t',ctx = Pretyping.understand env Evd.empty new_t in
let new_env = Environ.push_rel (n,None,t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -937,18 +903,18 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(depth + 1) b
in
mkGProd(n,new_t,new_b),
- Idset.filter not_free_in_t id_to_exclude
+ Id.Set.filter not_free_in_t id_to_exclude
| _ -> (* the first args is the name of the function! *)
assert false
end
- | GApp(loc1,GRef(loc2,eq_as_ref),[ty;GVar(loc3,id);rt])
- when eq_as_ref = Lazy.force Coqlib.coq_eq_ref && n = Anonymous
+ | GApp(loc1,GRef(loc2,eq_as_ref,_),[ty;GVar(loc3,id);rt])
+ when Globnames.eq_gr eq_as_ref (Lazy.force Coqlib.coq_eq_ref) && n == Anonymous
->
begin
try
observe (str "computing new type for eq : " ++ pr_glob_constr rt);
let t' =
- try Pretyping.Default.understand Evd.empty env t
+ try fst (Pretyping.understand env Evd.empty t)(*FIXME*)
with e when Errors.noncritical e -> raise Continue
in
let is_in_b = is_free_in id b in
@@ -970,36 +936,36 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
in
mkGProd(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 jmeq = Globnames.IndRef (fst (destInd (jmeq ()))) in
+ let ty',ctx = Pretyping.understand env Evd.empty ty in
let ind,args' = Inductive.find_inductive env ty' in
- let mib,_ = Global.lookup_inductive ind in
+ let mib,_ = Global.lookup_inductive (fst ind) in
let nparam = mib.Declarations.mind_nparams in
let params,arg' =
- ((Util.list_chop nparam args'))
+ ((Util.List.chop nparam args'))
in
let rt_typ =
- GApp(Util.dummy_loc,
- GRef (Util.dummy_loc,Libnames.IndRef ind),
+ GApp(Loc.ghost,
+ GRef (Loc.ghost,Globnames.IndRef (fst ind),None),
(List.map
(fun p -> Detyping.detype false []
- (Termops.names_of_rel_context env)
+ env Evd.empty
p) params)@(Array.to_list
(Array.make
(List.length args' - nparam)
(mkGHole ()))))
in
let eq' =
- GApp(loc1,GRef(loc2,jmeq),[ty;GVar(loc3,id);rt_typ;rt])
+ GApp(loc1,GRef(loc2,jmeq,None),[ty;GVar(loc3,id);rt_typ;rt])
in
observe (str "computing new type for jmeq : " ++ pr_glob_constr eq');
- let eq'_as_constr = Pretyping.Default.understand Evd.empty env eq' in
+ let eq'_as_constr,ctx = Pretyping.understand env Evd.empty 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
+ let ty' = snd (Util.List.chop nparam ty) in
List.fold_left2
(fun acc var_as_constr arg ->
if isRel var_as_constr
@@ -1011,11 +977,13 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
| Anonymous -> acc
| Name id' ->
(id',Detyping.detype false []
- (Termops.names_of_rel_context env)
+ env
+ Evd.empty
arg)::acc
else if isVar var_as_constr
then (destVar var_as_constr,Detyping.detype false []
- (Termops.names_of_rel_context env)
+ env
+ Evd.empty
arg)::acc
else acc
)
@@ -1041,7 +1009,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
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
+ let t',ctx = Pretyping.understand env Evd.empty eq' in
Environ.push_rel (n,None,t') env
in
let new_b,id_to_exclude =
@@ -1056,10 +1024,10 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(* J.F:. keep this comment it explain how to remove some meaningless equalities
if keep_eq then
mkGProd(n,t,new_b),id_to_exclude
- else new_b, Idset.add id id_to_exclude
+ else new_b, Id.Set.add id id_to_exclude
*)
- | GApp(loc1,GRef(loc2,eq_as_ref),[ty;rt1;rt2])
- when eq_as_ref = Lazy.force Coqlib.coq_eq_ref && n = Anonymous
+ | GApp(loc1,GRef(loc2,eq_as_ref,_),[ty;rt1;rt2])
+ when Globnames.eq_gr eq_as_ref (Lazy.force Coqlib.coq_eq_ref) && n == Anonymous
->
begin
try
@@ -1079,7 +1047,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
else raise Continue
with Continue ->
observe (str "computing new type for prod : " ++ pr_glob_constr rt);
- let t' = Pretyping.Default.understand Evd.empty env t in
+ let t',ctx = Pretyping.understand env Evd.empty t in
let new_env = Environ.push_rel (n,None,t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -1088,14 +1056,14 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(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)
- | _ -> mkGProd(n,t,new_b),Idset.filter not_free_in_t id_to_exclude
+ | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args ->
+ new_b,Id.Set.remove id
+ (Id.Set.filter not_free_in_t id_to_exclude)
+ | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude
end
| _ ->
observe (str "computing new type for prod : " ++ pr_glob_constr rt);
- let t' = Pretyping.Default.understand Evd.empty env t in
+ let t',ctx = Pretyping.understand env Evd.empty t in
let new_env = Environ.push_rel (n,None,t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -1104,17 +1072,17 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(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)
- | _ -> mkGProd(n,t,new_b),Idset.filter not_free_in_t id_to_exclude
+ | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args ->
+ new_b,Id.Set.remove id
+ (Id.Set.filter not_free_in_t id_to_exclude)
+ | _ -> mkGProd(n,t,new_b),Id.Set.filter not_free_in_t id_to_exclude
end
| GLambda(_,n,k,t,b) ->
begin
let not_free_in_t id = not (is_free_in id t) in
let new_crossed_types = t :: crossed_types in
observe (str "computing new type for lambda : " ++ pr_glob_constr rt);
- let t' = Pretyping.Default.understand Evd.empty env t in
+ let t',ctx = Pretyping.understand env Evd.empty t in
match n with
| Name id ->
let new_env = Environ.push_rel (n,None,t') env in
@@ -1124,19 +1092,19 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(args@[mkGVar id])new_crossed_types
(depth + 1 ) b
in
- if Idset.mem id id_to_exclude && depth >= nb_args
+ if Id.Set.mem id id_to_exclude && depth >= nb_args
then
- new_b, Idset.remove id (Idset.filter not_free_in_t id_to_exclude)
+ new_b, Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude)
else
- GProd(dummy_loc,n,k,t,new_b),Idset.filter not_free_in_t id_to_exclude
- | _ -> anomaly "Should not have an anonymous function here"
+ GProd(Loc.ghost,n,k,t,new_b),Id.Set.filter not_free_in_t id_to_exclude
+ | _ -> anomaly (Pp.str "Should not have an anonymous function here")
(* We have renamed all the anonymous functions during alpha_renaming phase *)
end
| GLetIn(_,n,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 t',ctx = Pretyping.understand env Evd.empty 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 =
@@ -1145,13 +1113,13 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
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)
- | _ -> GLetIn(dummy_loc,n,t,new_b),
- Idset.filter not_free_in_t id_to_exclude
+ | Name id when Id.Set.mem id id_to_exclude && depth >= nb_args ->
+ new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude)
+ | _ -> GLetIn(Loc.ghost,n,t,new_b),
+ Id.Set.filter not_free_in_t id_to_exclude
end
| GLetTuple(_,nal,(na,rto),t,b) ->
- assert (rto=None);
+ assert (Option.is_empty rto);
begin
let not_free_in_t id = not (is_free_in id t) in
let new_t,id_to_exclude' =
@@ -1161,7 +1129,7 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
args (crossed_types)
depth t
in
- let t' = Pretyping.Default.understand Evd.empty env new_t in
+ let t',ctx = Pretyping.understand env Evd.empty new_t in
let new_env = Environ.push_rel (na,None,t') env in
let new_b,id_to_exclude =
rebuild_cons new_env
@@ -1170,15 +1138,15 @@ let rec rebuild_cons env nb_args relname args crossed_types depth rt =
(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) *)
+(* | Name id when Id.Set.mem id id_to_exclude -> *)
+(* new_b,Id.Set.remove id (Id.Set.filter not_free_in_t id_to_exclude) *)
(* | _ -> *)
- GLetTuple(dummy_loc,nal,(na,None),t,new_b),
- Idset.filter not_free_in_t (Idset.union id_to_exclude id_to_exclude')
+ GLetTuple(Loc.ghost,nal,(na,None),t,new_b),
+ Id.Set.filter not_free_in_t (Id.Set.union id_to_exclude id_to_exclude')
end
- | _ -> mkGApp(mkGVar relname,args@[rt]),Idset.empty
+ | _ -> mkGApp(mkGVar relname,args@[rt]),Id.Set.empty
(* debuging wrapper *)
@@ -1201,7 +1169,7 @@ let rebuild_cons env nb_args relname args crossed_types rt =
*)
let rec compute_cst_params relnames params = function
| GRef _ | GVar _ | GEvar _ | GPatVar _ -> params
- | GApp(_,GVar(_,relname'),rtl) when Idset.mem relname' relnames ->
+ | GApp(_,GVar(_,relname'),rtl) when Id.Set.mem relname' relnames ->
compute_cst_params_from_app [] (params,rtl)
| GApp(_,f,args) ->
List.fold_left (compute_cst_params relnames) params (f::args)
@@ -1219,11 +1187,11 @@ 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',(GVar(_,id'))::rtl'
- when id_ord id id' == 0 && not is_defined ->
+ when Id.compare 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 * Glob_term.glob_constr * bool) list array) csts =
+let compute_params_name relnames (args : (Name.t * Glob_term.glob_constr * bool) list array) csts =
let rels_params =
Array.mapi
(fun i args ->
@@ -1237,12 +1205,12 @@ let compute_params_name relnames (args : (Names.name * Glob_term.glob_constr * b
let l = ref [] in
let _ =
try
- list_iter_i
+ List.iteri
(fun i ((n,nt,is_defined) as param) ->
- if array_for_all
+ if Array.for_all
(fun l ->
let (n',nt',is_defined') = List.nth l i in
- n = n' && Topconstr.eq_glob_constr nt nt' && is_defined = is_defined')
+ Name.equal n n' && Notation_ops.eq_glob_constr nt nt' && (is_defined : bool) == is_defined')
rels_params
then
l := param::!l
@@ -1255,22 +1223,23 @@ let compute_params_name relnames (args : (Names.name * Glob_term.glob_constr * b
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,GType None))
+ | Constrexpr.CProdN(loc,n,t') ->
+ Constrexpr.CProdN(loc,n,rebuild_return_type t')
+ | Constrexpr.CLetIn(loc,na,t,t') ->
+ Constrexpr.CLetIn(loc,na,t,rebuild_return_type t')
+ | _ -> Constrexpr.CProdN(Loc.ghost,[[Loc.ghost,Anonymous],
+ Constrexpr.Default Decl_kinds.Explicit,rt],
+ Constrexpr.CSort(Loc.ghost,GType []))
let do_build_inductive
- funnames (funsargs: (Names.name * glob_constr * bool) list list)
- returned_types
- (rtl:glob_constr list) =
+ mp_dp
+ funnames (funsargs: (Name.t * glob_constr * bool) list list)
+ returned_types
+ (rtl:glob_constr list) =
let _time1 = System.get_time () in
-(* Pp.msgnl (prlist_with_sep fnl Printer.pr_glob_constr rtl); *)
- let funnames_as_set = List.fold_right Idset.add funnames Idset.empty in
+ (* Pp.msgnl (prlist_with_sep fnl Printer.pr_glob_constr rtl); *)
+ let funnames_as_set = List.fold_right Id.Set.add funnames Id.Set.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
@@ -1281,12 +1250,22 @@ let do_build_inductive
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
+ let relnames_as_set = Array.fold_right Id.Set.add relnames Id.Set.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 (Constrintern.global_reference id)) env
+ let c =
+ match mp_dp with
+ | None -> (Constrintern.global_reference id)
+ | Some(mp,dp) -> mkConst (make_con mp dp (Label.of_id id))
+ in
+ Environ.push_named (id,None,
+ try
+ Typing.type_of env Evd.empty c
+ with Not_found ->
+ raise (UserError("do_build_inductive", str "Cannot handle partial fixpoint"))
+ ) env
)
funnames
(Global.env ())
@@ -1294,19 +1273,19 @@ let do_build_inductive
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 * Glob_term.glob_constr * bool ) list =
+ let rel_first_args :(Name.t * Glob_term.glob_constr * 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_glob_constr Idset.empty t,
+ Constrexpr.CLetIn(Loc.ghost,(Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t,
acc)
else
- Topconstr.CProdN
- (dummy_loc,
- [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_glob_constr Idset.empty t],
+ Constrexpr.CProdN
+ (Loc.ghost,
+ [[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t],
acc
)
)
@@ -1318,8 +1297,9 @@ let do_build_inductive
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
+ Util.Array.fold_left2 (fun env rel_name rel_ar ->
+ Environ.push_named (rel_name,None,
+ fst (with_full_print (Constrintern.interp_constr env Evd.empty) rel_ar)) env) env relnames rel_arities
in
(* and of the real constructors*)
let constr i res =
@@ -1344,9 +1324,9 @@ let do_build_inductive
(*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))
+ Id.of_string ((Id.to_string (mk_rel_id funnames.(i)))^"_"^(string_of_int !next_constructor_id))
in
- let rel_constructors i rt : (identifier*glob_constr) list =
+ let rel_constructors i rt : (Id.t*glob_constr) list =
next_constructor_id := (-1);
List.map (fun constr -> (mk_constructor_id i),constr) (constr i rt)
in
@@ -1360,19 +1340,19 @@ let do_build_inductive
rel_constructors
in
let rel_arity i funargs = (* Reduilding arities (with parameters) *)
- let rel_first_args :(Names.name * Glob_term.glob_constr * bool ) list =
- (snd (list_chop nrel_params funargs))
+ let rel_first_args :(Name.t * Glob_term.glob_constr * 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_glob_constr Idset.empty t,
+ Constrexpr.CLetIn(Loc.ghost,(Loc.ghost, n),with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t,
acc)
else
- Topconstr.CProdN
- (dummy_loc,
- [[(dummy_loc,n)],Topconstr.default_binder_kind,Constrextern.extern_glob_constr Idset.empty t],
+ Constrexpr.CProdN
+ (Loc.ghost,
+ [[(Loc.ghost,n)],Constrexpr_ops.default_binder_kind,with_full_print (Constrextern.extern_glob_constr Id.Set.empty) t],
acc
)
)
@@ -1384,31 +1364,40 @@ let do_build_inductive
Then save the graphs and reset Printing options to their primitive values
*)
let rel_arities = Array.mapi rel_arity funsargs in
+ let rel_params_ids =
+ List.fold_left
+ (fun acc (na,_,_) ->
+ match na with
+ Anonymous -> acc
+ | Name id -> id::acc
+ )
+ []
+ rels_params
+ in
let rel_params =
List.map
(fun (n,t,is_defined) ->
if is_defined
then
- Topconstr.LocalRawDef((dummy_loc,n), Constrextern.extern_glob_constr Idset.empty t)
+ Constrexpr.LocalRawDef((Loc.ghost,n), Constrextern.extern_glob_constr Id.Set.empty t)
else
- Topconstr.LocalRawAssum
- ([(dummy_loc,n)], Topconstr.default_binder_kind, Constrextern.extern_glob_constr Idset.empty t)
+ Constrexpr.LocalRawAssum
+ ([(Loc.ghost,n)], Constrexpr_ops.default_binder_kind, Constrextern.extern_glob_constr Id.Set.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_glob_type Idset.empty) ((* zeta_normalize *) t)
+ false,((Loc.ghost,id),
+ with_full_print
+ (Constrextern.extern_glob_type Id.Set.empty) ((* zeta_normalize *) (alpha_rt rel_params_ids t))
)
))
(rel_constructors)
in
let rel_ind i ext_rel_constructors =
- ((dummy_loc,relnames.(i)),
+ ((Loc.ghost,relnames.(i)),
rel_params,
Some rel_arities.(i),
ext_rel_constructors),[]
@@ -1437,7 +1426,7 @@ let do_build_inductive
(* in *)
let _time2 = System.get_time () in
try
- with_full_print (Flags.silently (Command.do_mutual_inductive rel_inds)) true
+ with_full_print (Flags.silently (Command.do_mutual_inductive rel_inds false false)) Decl_kinds.Finite
with
| UserError(s,msg) as e ->
let _time3 = System.get_time () in
@@ -1448,7 +1437,7 @@ let do_build_inductive
in
let msg =
str "while trying to define"++ spc () ++
- Ppvernac.pr_vernac (Vernacexpr.VernacInductive(Decl_kinds.Finite,false,repacked_rel_inds))
+ Ppvernac.pr_vernac (Vernacexpr.VernacInductive(false,Decl_kinds.Finite,repacked_rel_inds))
++ fnl () ++
msg
in
@@ -1463,7 +1452,7 @@ let do_build_inductive
in
let msg =
str "while trying to define"++ spc () ++
- Ppvernac.pr_vernac (Vernacexpr.VernacInductive(Decl_kinds.Finite,false,repacked_rel_inds))
+ Ppvernac.pr_vernac (Vernacexpr.VernacInductive(false,Decl_kinds.Finite,repacked_rel_inds))
++ fnl () ++
Errors.print reraise
in
@@ -1472,9 +1461,9 @@ let do_build_inductive
-let build_inductive funnames funsargs returned_types rtl =
+let build_inductive mp_dp funnames funsargs returned_types rtl =
try
- do_build_inductive funnames funsargs returned_types rtl
+ do_build_inductive mp_dp funnames funsargs returned_types rtl
with e when Errors.noncritical e -> raise (Building_graph e)
diff --git a/plugins/funind/glob_term_to_relation.mli b/plugins/funind/glob_term_to_relation.mli
index 5c91292b..b0a05ec3 100644
--- a/plugins/funind/glob_term_to_relation.mli
+++ b/plugins/funind/glob_term_to_relation.mli
@@ -1,5 +1,4 @@
-
-
+open Names
(*
[build_inductive parametrize funnames funargs returned_types bodies]
@@ -8,9 +7,10 @@
*)
val build_inductive :
- Names.identifier list -> (* The list of function name *)
- (Names.name*Glob_term.glob_constr*bool) list list -> (* The list of function args *)
- Topconstr.constr_expr list -> (* The list of function returned type *)
+ (ModPath.t * DirPath.t) option ->
+ Id.t list -> (* The list of function name *)
+ (Name.t*Glob_term.glob_constr*bool) list list -> (* The list of function args *)
+ Constrexpr.constr_expr list -> (* The list of function returned type *)
Glob_term.glob_constr list -> (* the list of body *)
unit
diff --git a/plugins/funind/glob_termops.ml b/plugins/funind/glob_termops.ml
index 6cc932b1..291f835e 100644
--- a/plugins/funind/glob_termops.ml
+++ b/plugins/funind/glob_termops.ml
@@ -1,24 +1,25 @@
open Pp
open Glob_term
+open Errors
open Util
open Names
-(* Ocaml 3.06 Map.S does not handle is_empty *)
-let idmap_is_empty m = m = Idmap.empty
+open Decl_kinds
+open Misctypes
(*
Some basic functions to rebuild glob_constr
- In each of them the location is Util.dummy_loc
+ In each of them the location is Loc.ghost
*)
-let mkGRef ref = GRef(dummy_loc,ref)
-let mkGVar id = GVar(dummy_loc,id)
-let mkGApp(rt,rtl) = GApp(dummy_loc,rt,rtl)
-let mkGLambda(n,t,b) = GLambda(dummy_loc,n,Explicit,t,b)
-let mkGProd(n,t,b) = GProd(dummy_loc,n,Explicit,t,b)
-let mkGLetIn(n,t,b) = GLetIn(dummy_loc,n,t,b)
-let mkGCases(rto,l,brl) = GCases(dummy_loc,Term.RegularStyle,rto,l,brl)
-let mkGSort s = GSort(dummy_loc,s)
-let mkGHole () = GHole(dummy_loc,Evd.BinderType Anonymous)
-let mkGCast(b,t) = GCast(dummy_loc,b,CastConv (Term.DEFAULTcast,t))
+let mkGRef ref = GRef(Loc.ghost,ref,None)
+let mkGVar id = GVar(Loc.ghost,id)
+let mkGApp(rt,rtl) = GApp(Loc.ghost,rt,rtl)
+let mkGLambda(n,t,b) = GLambda(Loc.ghost,n,Explicit,t,b)
+let mkGProd(n,t,b) = GProd(Loc.ghost,n,Explicit,t,b)
+let mkGLetIn(n,t,b) = GLetIn(Loc.ghost,n,t,b)
+let mkGCases(rto,l,brl) = GCases(Loc.ghost,Term.RegularStyle,rto,l,brl)
+let mkGSort s = GSort(Loc.ghost,s)
+let mkGHole () = GHole(Loc.ghost,Evar_kinds.BinderType Anonymous,Misctypes.IntroAnonymous,None)
+let mkGCast(b,t) = GCast(Loc.ghost,b,CastConv t)
(*
Some basic functions to decompose glob_constrs
@@ -107,7 +108,7 @@ let glob_make_or t1 t2 = mkGApp (mkGRef(Lazy.force Coqlib.coq_or_ref),[t1;t2])
to [P1 \/ ( .... \/ Pn)]
*)
let rec glob_make_or_list = function
- | [] -> raise (Invalid_argument "mk_or")
+ | [] -> invalid_arg "mk_or"
| [e] -> e
| e::l -> glob_make_or e (glob_make_or_list l)
@@ -115,7 +116,7 @@ let rec glob_make_or_list = function
let remove_name_from_mapping mapping na =
match na with
| Anonymous -> mapping
- | Name id -> Idmap.remove id mapping
+ | Name id -> Id.Map.remove id mapping
let change_vars =
let rec change_vars mapping rt =
@@ -124,7 +125,7 @@ let change_vars =
| GVar(loc,id) ->
let new_id =
try
- Idmap.find id mapping
+ Id.Map.find id mapping
with Not_found -> id
in
GVar(loc,new_id)
@@ -179,13 +180,12 @@ let change_vars =
| GRec _ -> error "Local (co)fixes are not supported"
| GSort _ -> rt
| GHole _ -> rt
- | GCast(loc,b,CastConv (k,t)) ->
- GCast(loc,change_vars mapping b, CastConv (k,change_vars mapping t))
- | GCast(loc,b,CastCoerce) ->
- GCast(loc,change_vars mapping b,CastCoerce)
+ | GCast(loc,b,c) ->
+ GCast(loc,change_vars mapping b,
+ Miscops.map_cast_type (change_vars mapping) c)
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
+ let new_mapping = List.fold_right Id.Map.remove idl mapping in
+ if Id.Map.is_empty new_mapping
then br
else (loc,idl,patl,change_vars new_mapping res)
in
@@ -197,27 +197,27 @@ 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 new_id),(new_id::excluded),Id.Map.empty
| PatVar(loc,Name id) ->
- if List.mem id excluded
+ if Id.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
+ (Id.Map.add id new_id Id.Map.empty)
+ else pat,excluded,Id.Map.empty
| PatCstr(loc,constr,patl,na) ->
let new_na,new_excluded,map =
match na with
- | Name id when List.mem id excluded ->
+ | Name id when Id.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
+ Name new_id,new_id::excluded, Id.Map.add id new_id Id.Map.empty
+ | _ -> na,excluded,Id.Map.empty
in
let new_patl,new_excluded,new_map =
List.fold_left
(fun (patl,excluded,map) pat ->
let new_pat,new_excluded,new_map = alpha_pat excluded pat in
- (new_pat::patl,new_excluded,Idmap.fold Idmap.add new_map map)
+ (new_pat::patl,new_excluded,Id.Map.fold Id.Map.add new_map map)
)
([],new_excluded,map)
patl
@@ -229,9 +229,9 @@ let alpha_patl excluded patl =
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_pat::patl,new_excluded,(Id.Map.fold Id.Map.add new_map map)
)
- ([],excluded,Idmap.empty)
+ ([],excluded,Id.Map.empty)
patl
in
(List.rev patl,new_excluded,map)
@@ -263,7 +263,7 @@ let rec alpha_rt excluded rt =
match rt with
| GRef _ | GVar _ | GEvar _ | GPatVar _ -> rt
| GLambda(loc,Anonymous,k,t,b) ->
- let new_id = Namegen.next_ident_away (id_of_string "_x") excluded in
+ 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
@@ -279,10 +279,10 @@ let rec alpha_rt excluded rt =
| GLambda(loc,Name id,k,t,b) ->
let new_id = Namegen.next_ident_away id excluded in
let t,b =
- if new_id = id
+ if Id.equal new_id id
then t,b
else
- let replace = change_vars (Idmap.add id new_id Idmap.empty) in
+ let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in
(t,replace b)
in
let new_excluded = new_id::excluded in
@@ -293,10 +293,10 @@ let rec alpha_rt excluded rt =
let new_id = Namegen.next_ident_away id excluded in
let new_excluded = new_id::excluded in
let t,b =
- if new_id = id
+ if Id.equal new_id id
then t,b
else
- let replace = change_vars (Idmap.add id new_id Idmap.empty) in
+ let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in
(t,replace b)
in
let new_t = alpha_rt new_excluded t in
@@ -305,10 +305,10 @@ let rec alpha_rt excluded rt =
| GLetIn(loc,Name id,t,b) ->
let new_id = Namegen.next_ident_away id excluded in
let t,b =
- if new_id = id
+ if Id.equal new_id id
then t,b
else
- let replace = change_vars (Idmap.add id new_id Idmap.empty) in
+ let replace = change_vars (Id.Map.add id new_id Id.Map.empty) in
(t,replace b)
in
let new_excluded = new_id::excluded in
@@ -325,18 +325,18 @@ let rec alpha_rt excluded rt =
| Anonymous -> (na::nal,excluded,mapping)
| Name id ->
let new_id = Namegen.next_ident_away id excluded in
- if new_id = id
+ if Id.equal new_id id
then
na::nal,id::excluded,mapping
else
- (Name new_id)::nal,id::excluded,(Idmap.add id new_id mapping)
+ (Name new_id)::nal,id::excluded,(Id.Map.add id new_id mapping)
)
- ([],excluded,Idmap.empty)
+ ([],excluded,Id.Map.empty)
nal
in
let new_nal = List.rev rev_new_nal in
let new_rto,new_t,new_b =
- if idmap_is_empty mapping
+ if Id.Map.is_empty mapping
then rto,t,b
else let replace = change_vars mapping in
(Option.map replace rto, t,replace b)
@@ -359,10 +359,9 @@ let rec alpha_rt excluded rt =
| GRec _ -> error "Not handled GRec"
| GSort _ -> rt
| GHole _ -> rt
- | GCast (loc,b,CastConv (k,t)) ->
- GCast(loc,alpha_rt excluded b,CastConv(k,alpha_rt excluded t))
- | GCast (loc,b,CastCoerce) ->
- GCast(loc,alpha_rt excluded b,CastCoerce)
+ | GCast (loc,b,c) ->
+ GCast(loc,alpha_rt excluded b,
+ Miscops.map_cast_type (alpha_rt excluded) c)
| GApp(loc,f,args) ->
GApp(loc,
alpha_rt excluded f,
@@ -385,14 +384,14 @@ and alpha_br excluded (loc,ids,patl,res) =
let is_free_in id =
let rec is_free_in = function
| GRef _ -> false
- | GVar(_,id') -> id_ord id' id == 0
+ | GVar(_,id') -> Id.compare id' id == 0
| GEvar _ -> false
| GPatVar _ -> false
| GApp(_,rt,rtl) -> List.exists is_free_in (rt::rtl)
| GLambda(_,n,_,t,b) | GProd(_,n,_,t,b) | GLetIn(_,n,t,b) ->
let check_in_b =
match n with
- | Name id' -> id_ord id' id <> 0
+ | Name id' -> not (Id.equal id' id)
| _ -> true
in
is_free_in t || (check_in_b && is_free_in b)
@@ -401,7 +400,7 @@ let is_free_in id =
List.exists is_free_in_br brl
| GLetTuple(_,nal,_,b,t) ->
let check_in_nal =
- not (List.exists (function Name id' -> id'= id | _ -> false) nal)
+ not (List.exists (function Name id' -> Id.equal id' id | _ -> false) nal)
in
is_free_in t || (check_in_nal && is_free_in b)
@@ -410,10 +409,10 @@ let is_free_in id =
| GRec _ -> raise (UserError("",str "Not handled GRec"))
| GSort _ -> false
| GHole _ -> false
- | GCast (_,b,CastConv (_,t)) -> is_free_in b || is_free_in t
+ | GCast (_,b,(CastConv t|CastVM t|CastNative t)) -> is_free_in b || is_free_in t
| GCast (_,b,CastCoerce) -> is_free_in b
and is_free_in_br (_,ids,_,rt) =
- (not (List.mem id ids)) && is_free_in rt
+ (not (Id.List.mem id ids)) && is_free_in rt
in
is_free_in
@@ -425,7 +424,7 @@ let rec pattern_to_term = function
mkGVar id
| PatCstr(loc,constr,patternl,_) ->
let cst_narg =
- Inductiveops.mis_constructor_nargs_env
+ Inductiveops.constructor_nallargs_env
(Global.env ())
constr
in
@@ -439,7 +438,7 @@ let rec pattern_to_term = function
let patl_as_term =
List.map pattern_to_term patternl
in
- mkGApp(mkGRef(Libnames.ConstructRef constr),
+ mkGApp(mkGRef(Globnames.ConstructRef constr),
implicit_args@patl_as_term
)
@@ -449,7 +448,7 @@ let replace_var_by_term x_id term =
let rec replace_var_by_pattern rt =
match rt with
| GRef _ -> rt
- | GVar(_,id) when id_ord id x_id == 0 -> term
+ | GVar(_,id) when Id.compare id x_id == 0 -> term
| GVar _ -> rt
| GEvar _ -> rt
| GPatVar _ -> rt
@@ -458,7 +457,7 @@ let replace_var_by_term x_id term =
replace_var_by_pattern rt',
List.map replace_var_by_pattern rtl
)
- | GLambda(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt
+ | GLambda(_,Name id,_,_,_) when Id.compare id x_id == 0 -> rt
| GLambda(loc,name,k,t,b) ->
GLambda(loc,
name,
@@ -466,7 +465,7 @@ let replace_var_by_term x_id term =
replace_var_by_pattern t,
replace_var_by_pattern b
)
- | GProd(_,Name id,_,_,_) when id_ord id x_id == 0 -> rt
+ | GProd(_,Name id,_,_,_) when Id.compare id x_id == 0 -> rt
| GProd(loc,name,k,t,b) ->
GProd(loc,
name,
@@ -474,7 +473,7 @@ let replace_var_by_term x_id term =
replace_var_by_pattern t,
replace_var_by_pattern b
)
- | GLetIn(_,Name id,_,_) when id_ord id x_id == 0 -> rt
+ | GLetIn(_,Name id,_,_) when Id.compare id x_id == 0 -> rt
| GLetIn(loc,name,def,b) ->
GLetIn(loc,
name,
@@ -482,7 +481,7 @@ let replace_var_by_term x_id term =
replace_var_by_pattern b
)
| GLetTuple(_,nal,_,_,_)
- when List.exists (function Name id -> id = x_id | _ -> false) nal ->
+ when List.exists (function Name id -> Id.equal id x_id | _ -> false) nal ->
rt
| GLetTuple(loc,nal,(na,rto),def,b) ->
GLetTuple(loc,
@@ -506,12 +505,11 @@ let replace_var_by_term x_id term =
| GRec _ -> raise (UserError("",str "Not handled GRec"))
| GSort _ -> rt
| GHole _ -> rt
- | GCast(loc,b,CastConv(k,t)) ->
- GCast(loc,replace_var_by_pattern b,CastConv(k,replace_var_by_pattern t))
- | GCast(loc,b,CastCoerce) ->
- GCast(loc,replace_var_by_pattern b,CastCoerce)
+ | GCast(loc,b,c) ->
+ GCast(loc,replace_var_by_pattern b,
+ Miscops.map_cast_type replace_var_by_pattern c)
and replace_var_by_pattern_br ((loc,idl,patl,res) as br) =
- if List.exists (fun id -> id_ord id x_id == 0) idl
+ if List.exists (fun id -> Id.compare id x_id == 0) idl
then br
else (loc,idl,patl,replace_var_by_pattern res)
in
@@ -529,13 +527,12 @@ let rec are_unifiable_aux = function
match eq with
| PatVar _,_ | _,PatVar _ -> are_unifiable_aux eqs
| PatCstr(_,constructor1,cpl1,_),PatCstr(_,constructor2,cpl2,_) ->
- if constructor2 <> constructor1
+ if not (eq_constructor constructor2 constructor1)
then raise NotUnifiable
else
let eqs' =
- try ((List.combine cpl1 cpl2)@eqs)
- with e when Errors.noncritical e ->
- anomaly "are_unifiable_aux"
+ try (List.combine cpl1 cpl2) @ eqs
+ with Invalid_argument _ -> anomaly (Pp.str "are_unifiable_aux")
in
are_unifiable_aux eqs'
@@ -552,13 +549,12 @@ let rec eq_cases_pattern_aux = function
match eq with
| PatVar _,PatVar _ -> eq_cases_pattern_aux eqs
| PatCstr(_,constructor1,cpl1,_),PatCstr(_,constructor2,cpl2,_) ->
- if constructor2 <> constructor1
+ if not (eq_constructor constructor2 constructor1)
then raise NotUnifiable
else
let eqs' =
- try ((List.combine cpl1 cpl2)@eqs)
- with e when Errors.noncritical e ->
- anomaly "eq_cases_pattern_aux"
+ try (List.combine cpl1 cpl2) @ eqs
+ with Invalid_argument _ -> anomaly (Pp.str "eq_cases_pattern_aux")
in
eq_cases_pattern_aux eqs'
| _ -> raise NotUnifiable
@@ -574,13 +570,13 @@ let eq_cases_pattern pat1 pat2 =
let ids_of_pat =
let rec ids_of_pat ids = function
| PatVar(_,Anonymous) -> ids
- | PatVar(_,Name id) -> Idset.add id ids
+ | PatVar(_,Name id) -> Id.Set.add id ids
| PatCstr(_,_,patl,_) -> List.fold_left ids_of_pat ids patl
in
- ids_of_pat Idset.empty
+ ids_of_pat Id.Set.empty
let id_of_name = function
- | Names.Anonymous -> id_of_string "x"
+ | Names.Anonymous -> Id.of_string "x"
| Names.Name x -> x
(* TODO: finish Rec caes *)
@@ -594,7 +590,7 @@ let ids_of_glob_constr c =
| GLambda (loc,na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc
| GProd (loc,na,k,ty,c) -> idof na :: ids_of_glob_constr [] ty @ ids_of_glob_constr [] c @ acc
| GLetIn (loc,na,b,c) -> idof na :: ids_of_glob_constr [] b @ ids_of_glob_constr [] c @ acc
- | GCast (loc,c,CastConv(k,t)) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] t @ acc
+ | GCast (loc,c,(CastConv t|CastVM t|CastNative t)) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] t @ acc
| GCast (loc,c,CastCoerce) -> ids_of_glob_constr [] c @ acc
| GIf (loc,c,(na,po),b1,b2) -> ids_of_glob_constr [] c @ ids_of_glob_constr [] b1 @ ids_of_glob_constr [] b2 @ acc
| GLetTuple (_,nal,(na,po),b,c) ->
@@ -605,7 +601,7 @@ let ids_of_glob_constr c =
| (GSort _ | GHole _ | GRef _ | GEvar _ | GPatVar _) -> []
in
(* build the set *)
- List.fold_left (fun acc x -> Idset.add x acc) Idset.empty (ids_of_glob_constr [] c)
+ List.fold_left (fun acc x -> Id.Set.add x acc) Id.Set.empty (ids_of_glob_constr [] c)
@@ -662,10 +658,9 @@ let zeta_normalize =
| GRec _ -> raise (UserError("",str "Not handled GRec"))
| GSort _ -> rt
| GHole _ -> rt
- | GCast(loc,b,CastConv(k,t)) ->
- GCast(loc,zeta_normalize_term b,CastConv(k,zeta_normalize_term t))
- | GCast(loc,b,CastCoerce) ->
- GCast(loc,zeta_normalize_term b,CastCoerce)
+ | GCast(loc,b,c) ->
+ GCast(loc,zeta_normalize_term b,
+ Miscops.map_cast_type zeta_normalize_term c)
and zeta_normalize_br (loc,idl,patl,res) =
(loc,idl,patl,zeta_normalize_term res)
in
@@ -680,7 +675,7 @@ let expand_as =
match pat with
| PatVar _ -> map
| PatCstr(_,_,patl,Name id) ->
- Idmap.add id (pattern_to_term pat) (List.fold_left add_as map patl)
+ Id.Map.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 =
@@ -689,7 +684,7 @@ let expand_as =
| GVar(_,id) ->
begin
try
- Idmap.find id map
+ Id.Map.find id map
with Not_found -> rt
end
| GApp(loc,f,args) -> GApp(loc,expand_as map f,List.map (expand_as map) args)
@@ -703,12 +698,13 @@ let expand_as =
GIf(loc,expand_as map e,(na,Option.map (expand_as map) po),
expand_as map br1, expand_as map br2)
| GRec _ -> error "Not handled GRec"
- | GCast(loc,b,CastConv(kind,t)) -> GCast(loc,expand_as map b,CastConv(kind,expand_as map t))
- | GCast(loc,b,CastCoerce) -> GCast(loc,expand_as map b,CastCoerce)
+ | GCast(loc,b,c) ->
+ GCast(loc,expand_as map b,
+ Miscops.map_cast_type (expand_as map) c)
| GCases(loc,sty,po,el,brl) ->
GCases(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
+ expand_as Id.Map.empty
diff --git a/plugins/funind/glob_termops.mli b/plugins/funind/glob_termops.mli
index bfd15357..0f10636f 100644
--- a/plugins/funind/glob_termops.mli
+++ b/plugins/funind/glob_termops.mli
@@ -1,11 +1,9 @@
+open Names
open Glob_term
-
-(* Ocaml 3.06 Map.S does not handle is_empty *)
-val idmap_is_empty : 'a Names.Idmap.t -> bool
-
+open Misctypes
(* [get_pattern_id pat] returns a list of all the variable appearing in [pat] *)
-val get_pattern_id : cases_pattern -> Names.identifier list
+val get_pattern_id : cases_pattern -> Id.t list
(* [pattern_to_term pat] returns a glob_constr corresponding to [pat].
[pat] must not contain occurences of anonymous pattern
@@ -14,14 +12,14 @@ val pattern_to_term : cases_pattern -> glob_constr
(*
Some basic functions to rebuild glob_constr
- In each of them the location is Util.dummy_loc
+ In each of them the location is Util.Loc.ghost
*)
-val mkGRef : Libnames.global_reference -> glob_constr
-val mkGVar : Names.identifier -> glob_constr
+val mkGRef : Globnames.global_reference -> glob_constr
+val mkGVar : Id.t -> glob_constr
val mkGApp : glob_constr*(glob_constr list) -> glob_constr
-val mkGLambda : Names.name * glob_constr * glob_constr -> glob_constr
-val mkGProd : Names.name * glob_constr * glob_constr -> glob_constr
-val mkGLetIn : Names.name * glob_constr * glob_constr -> glob_constr
+val mkGLambda : Name.t * glob_constr * glob_constr -> glob_constr
+val mkGProd : Name.t * glob_constr * glob_constr -> glob_constr
+val mkGLetIn : Name.t * glob_constr * glob_constr -> glob_constr
val mkGCases : glob_constr option * tomatch_tuples * cases_clauses -> glob_constr
val mkGSort : glob_sort -> glob_constr
val mkGHole : unit -> glob_constr (* we only build Evd.BinderType Anonymous holes *)
@@ -30,15 +28,15 @@ val mkGCast : glob_constr* glob_constr -> glob_constr
Some basic functions to decompose glob_constrs
These are analogous to the ones constrs
*)
-val glob_decompose_prod : glob_constr -> (Names.name*glob_constr) list * glob_constr
+val glob_decompose_prod : glob_constr -> (Name.t*glob_constr) list * glob_constr
val glob_decompose_prod_or_letin :
- glob_constr -> (Names.name*glob_constr option*glob_constr option) list * glob_constr
-val glob_decompose_prod_n : int -> glob_constr -> (Names.name*glob_constr) list * glob_constr
+ glob_constr -> (Name.t*glob_constr option*glob_constr option) list * glob_constr
+val glob_decompose_prod_n : int -> glob_constr -> (Name.t*glob_constr) list * glob_constr
val glob_decompose_prod_or_letin_n : int -> glob_constr ->
- (Names.name*glob_constr option*glob_constr option) list * glob_constr
-val glob_compose_prod : glob_constr -> (Names.name*glob_constr) list -> glob_constr
+ (Name.t*glob_constr option*glob_constr option) list * glob_constr
+val glob_compose_prod : glob_constr -> (Name.t*glob_constr) list -> glob_constr
val glob_compose_prod_or_letin: glob_constr ->
- (Names.name*glob_constr option*glob_constr option) list -> glob_constr
+ (Name.t*glob_constr option*glob_constr option) list -> glob_constr
val glob_decompose_app : glob_constr -> glob_constr*(glob_constr list)
@@ -60,7 +58,7 @@ val glob_make_or_list : glob_constr list -> glob_constr
(* Replace the var mapped in the glob_constr/context *)
-val change_vars : Names.identifier Names.Idmap.t -> glob_constr -> glob_constr
+val change_vars : Id.t Id.Map.t -> glob_constr -> glob_constr
@@ -72,27 +70,27 @@ val change_vars : Names.identifier Names.Idmap.t -> glob_constr -> glob_constr
[avoid] with the variables appearing in the result.
*)
val alpha_pat :
- Names.Idmap.key list ->
+ Id.Map.key list ->
Glob_term.cases_pattern ->
- Glob_term.cases_pattern * Names.Idmap.key list *
- Names.identifier Names.Idmap.t
+ Glob_term.cases_pattern * Id.Map.key list *
+ Id.t Id.Map.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 -> glob_constr -> glob_constr
+val alpha_rt : Id.t list -> glob_constr -> glob_constr
(* same as alpha_rt but for case branches *)
-val alpha_br : Names.identifier list ->
- Util.loc * Names.identifier list * Glob_term.cases_pattern list *
+val alpha_br : Id.t list ->
+ Loc.t * Id.t list * Glob_term.cases_pattern list *
Glob_term.glob_constr ->
- Util.loc * Names.identifier list * Glob_term.cases_pattern list *
+ Loc.t * Id.t list * Glob_term.cases_pattern list *
Glob_term.glob_constr
(* Reduction function *)
val replace_var_by_term :
- Names.identifier ->
+ Id.t ->
Glob_term.glob_constr -> Glob_term.glob_constr -> Glob_term.glob_constr
@@ -100,7 +98,7 @@ val replace_var_by_term :
(*
[is_free_in id rt] checks if [id] is a free variable in [rt]
*)
-val is_free_in : Names.identifier -> glob_constr -> bool
+val is_free_in : Id.t -> glob_constr -> bool
val are_unifiable : cases_pattern -> cases_pattern -> bool
@@ -109,13 +107,13 @@ val eq_cases_pattern : cases_pattern -> cases_pattern -> bool
(*
- ids_of_pat : cases_pattern -> Idset.t
+ ids_of_pat : cases_pattern -> Id.Set.t
returns the set of variables appearing in a pattern
*)
-val ids_of_pat : cases_pattern -> Names.Idset.t
+val ids_of_pat : cases_pattern -> Id.Set.t
(* TODO: finish this function (Fix not treated) *)
-val ids_of_glob_constr: glob_constr -> Names.Idset.t
+val ids_of_glob_constr: glob_constr -> Id.Set.t
(*
removing let_in construction in a glob_constr
diff --git a/plugins/funind/indfun.ml b/plugins/funind/indfun.ml
index d2c065a0..6dbd61cf 100644
--- a/plugins/funind/indfun.ml
+++ b/plugins/funind/indfun.ml
@@ -1,11 +1,16 @@
+open Errors
open Util
open Names
open Term
open Pp
open Indfun_common
open Libnames
+open Globnames
open Glob_term
open Declarations
+open Declareops
+open Misctypes
+open Decl_kinds
let is_rec_info scheme_info =
let test_branche min acc (_,_,br) =
@@ -14,15 +19,13 @@ let is_rec_info scheme_info =
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
+ Int.Set.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)
+ 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
+ Tactics.induction_destruct (is_rec_info scheme_info) false
let functional_induction with_clean c princl pat =
Dumpglob.pause ();
@@ -33,7 +36,7 @@ let functional_induction with_clean c princl pat =
| None -> (* No principle is given let's find the good one *)
begin
match kind_of_term f with
- | Const c' ->
+ | Const (c',u) ->
let princ_option =
let finfo = (* we first try to find out a graph on f *)
try find_Function_infos c'
@@ -54,7 +57,7 @@ let functional_induction with_clean c princl pat =
(or f_rec, f_rect) i*)
let princ_name =
Indrec.make_elimination_ident
- (id_of_label (con_label c'))
+ (Label.to_id (con_label c'))
(Tacticals.elimination_sort_of_goal g)
in
try
@@ -63,7 +66,7 @@ let functional_induction with_clean c princl pat =
errorlabstrm "" (str "Cannot find induction principle for "
++Printer.pr_lconstr (mkConst c') )
in
- (princ,Glob_term.NoBindings, Tacmach.pf_type_of g princ)
+ (princ,NoBindings, Tacmach.pf_type_of g princ)
| _ -> raise (UserError("",str "functional induction must be used with a function" ))
end
| Some ((princ,binding)) ->
@@ -75,50 +78,43 @@ let functional_induction with_clean c princl pat =
if princ_infos.Tactics.farg_in_concl
then [c] else []
in
- List.map (fun c -> Tacexpr.ElimOnConstr (Evd.empty,(c,NoBindings))) (args@c_list)
+ let encoded_pat_as_patlist =
+ List.make (List.length args + List.length c_list - 1) None @ [pat] in
+ List.map2 (fun c pat -> ((None,Tacexpr.ElimOnConstr (fun env sigma -> sigma,(c,NoBindings))),(None,pat),None))
+ (args@c_list) encoded_pat_as_patlist
in
let princ' = Some (princ,bindings) in
let princ_vars =
List.fold_right
- (fun a acc ->
- try Idset.add (destVar a) acc
- with e when Errors.noncritical e -> acc
- )
+ (fun a acc -> try Id.Set.add (destVar a) acc with DestKO -> acc)
args
- Idset.empty
+ Id.Set.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 old_idl = List.fold_right Id.Set.add (Tacmach.pf_ids_of_hyps g) Id.Set.empty in
+ let old_idl = Id.Set.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
- )
+ List.filter (fun id -> not (Id.Set.mem id old_idl))
(Tacmach.pf_ids_of_hyps g)
in
let flag =
- Glob_term.Cbv
- {Glob_term.all_flags
- with Glob_term.rDelta = false;
+ Genredexpr.Cbv
+ {Redops.all_flags
+ with Genredexpr.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)
+ (Tacticals.tclMAP (fun id -> Tacticals.tclTRY (Proofview.V82.of_tactic (Equality.subst_gen (do_rewrite_dependent ()) [id]))) idl )
+ (Tactics.reduce flag Locusops.allHypsAndConcl)
g
else Tacticals.tclIDTAC g
in
Tacticals.tclTHEN
- (choose_dest_or_ind
+ (Proofview.V82.of_tactic (choose_dest_or_ind
princ_infos
- args_as_induction_constr
- princ'
- (None,pat)
- None)
+ (args_as_induction_constr,princ')))
subst_and_reduce
g
in
@@ -127,14 +123,14 @@ let functional_induction with_clean c princl pat =
let rec abstract_glob_constr c = function
| [] -> c
- | Topconstr.LocalRawDef (x,b)::bl -> Topconstr.mkLetInC(x,b,abstract_glob_constr c bl)
- | Topconstr.LocalRawAssum (idl,k,t)::bl ->
- List.fold_right (fun x b -> Topconstr.mkLambdaC([x],k,t,b)) idl
+ | Constrexpr.LocalRawDef (x,b)::bl -> Constrexpr_ops.mkLetInC(x,b,abstract_glob_constr c bl)
+ | Constrexpr.LocalRawAssum (idl,k,t)::bl ->
+ List.fold_right (fun x b -> Constrexpr_ops.mkLambdaC([x],k,t,b)) idl
(abstract_glob_constr c bl)
-let interp_casted_constr_with_implicits sigma env impls c =
- Constrintern.intern_gen false sigma env ~impls
- ~allow_patvar:false ~ltacvars:([],[]) c
+let interp_casted_constr_with_implicits env sigma impls c =
+ Constrintern.intern_gen Pretyping.WithoutTypeConstraint env ~impls
+ ~allow_patvar:false c
(*
Construct a fixpoint as a Glob_term
@@ -149,26 +145,21 @@ let build_newrecursive
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, Idmap.add recname impl impls))
+ let arityc = Constrexpr_ops.prod_constr_expr arityc bl in
+ let arity,ctx = Constrintern.interp_type env0 sigma arityc in
+ let evdref = ref (Evd.from_env env0) in
+ let _, (_, impls') = Constrintern.interp_context_evars env evdref bl in
+ let impl = Constrintern.compute_internalization_data env0 Constrintern.Recursive arity impls' in
+ (Environ.push_named (recname,None,arity) env, Id.Map.add recname impl impls))
(env0,Constrintern.empty_internalization_env) lnameargsardef in
let recdef =
(* Declare local notations *)
- let fs = States.freeze() in
- let def =
- try
- List.map
- (fun (_,bl,_,def) ->
- let def = abstract_glob_constr def bl in
- interp_casted_constr_with_implicits
- sigma rec_sign rec_impls def
- )
- lnameargsardef
- with reraise ->
- States.unfreeze fs; raise reraise in
- States.unfreeze fs; def
+ let f (_,bl,_,def) =
+ let def = abstract_glob_constr def bl in
+ interp_casted_constr_with_implicits
+ rec_sign sigma rec_impls def
+ in
+ States.with_state_protection (List.map f) lnameargsardef
in
recdef,rec_impls
@@ -178,15 +169,15 @@ let build_newrecursive l =
match body_opt with
| Some body ->
(fixna,bll,ar,body)
- | None -> user_err_loc (dummy_loc,"Function",str "Body of Function must be given")
+ | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given")
) l
in
build_newrecursive l'
(* 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 is_rec names =
+ let names = List.fold_right Id.Set.add names Id.Set.empty in
+ let check_id id names = Id.Set.mem id names in
let rec lookup names = function
| GVar(_,id) -> check_id id names
| GRef _ | GEvar _ | GPatVar _ | GSort _ | GHole _ -> false
@@ -195,11 +186,11 @@ let rec is_rec names =
| GIf(_,b,_,lhs,rhs) ->
(lookup names b) || (lookup names lhs) || (lookup names rhs)
| GLetIn(_,na,t,b) | GLambda(_,na,_,t,b) | GProd(_,na,_,t,b) ->
- lookup names t || lookup (Nameops.name_fold Idset.remove na names) b
+ lookup names t || lookup (Nameops.name_fold Id.Set.remove na names) b
| GLetTuple(_,nal,_,t,b) -> lookup names t ||
lookup
(List.fold_left
- (fun acc na -> Nameops.name_fold Idset.remove na acc)
+ (fun acc na -> Nameops.name_fold Id.Set.remove na acc)
names
nal
)
@@ -209,7 +200,7 @@ let rec is_rec names =
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
+ let new_names = List.fold_right Id.Set.remove idl names in
lookup new_names rt
in
lookup names
@@ -217,8 +208,8 @@ let rec is_rec 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
+ | Constrexpr.LocalRawDef _::bl -> 1 + local_binders_length bl
+ | Constrexpr.LocalRawAssum (idl,_,_)::bl -> List.length idl + local_binders_length bl
let prepare_body ((name,_,args,types,_),_) rt =
let n = local_binders_length args in
@@ -226,12 +217,14 @@ let prepare_body ((name,_,args,types,_),_) rt =
let fun_args,rt' = chop_rlambda_n n rt in
(fun_args,rt')
+let process_vernac_interp_error e =
+ fst (Cerrors.process_vernac_interp_error (e, Exninfo.null))
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 (Constrintern.global_reference id)) fix_names
+ List.map (fun id -> fst (destConst (Constrintern.global_reference id))) fix_names
in
(*
Then we check that the graphs have been defined
@@ -248,38 +241,45 @@ let derive_inversion fix_names =
Ensures by : register_built
i*)
(List.map
- (fun id -> destInd (Constrintern.global_reference (mk_rel_id id)))
+ (fun id -> fst (destInd (Constrintern.global_reference (mk_rel_id id))))
fix_names
)
with e when Errors.noncritical e ->
- let e' = Cerrors.process_vernac_interp_error e in
+ let e' = process_vernac_interp_error e in
msg_warning
(str "Cannot build inversion information" ++
if do_observe () then (fnl() ++ Errors.print e') else mt ())
with e when Errors.noncritical e -> ()
let warning_error names e =
- let e = Cerrors.process_vernac_interp_error e in
+ let e = process_vernac_interp_error e in
let e_explain e =
match e with
- | ToShow e -> spc () ++ Errors.print e
- | _ -> if do_observe () then (spc () ++ Errors.print e) else mt ()
+ | ToShow e ->
+ let e = process_vernac_interp_error e in
+ spc () ++ Errors.print e
+ | _ ->
+ if do_observe ()
+ then
+ let e = process_vernac_interp_error e in
+ (spc () ++ Errors.print 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)
+ 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)
+ Pp.msg_warning
+ (str "Cannot define principle(s) for "++
+ h 1 (prlist_with_sep (fun _ -> str","++spc ()) Ppconstr.pr_id names) ++
+ e_explain e)
| _ -> raise e
let error_error names e =
- let e = Cerrors.process_vernac_interp_error e in
+ let e = process_vernac_interp_error e in
let e_explain e =
match e with
| ToShow e -> spc () ++ Errors.print e
@@ -293,7 +293,7 @@ let error_error names e =
e_explain e)
| _ -> raise e
-let generate_principle on_error
+let generate_principle mp_dp on_error
is_general do_built (fix_rec_l:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) recdefs interactive_proof
(continue_proof : int -> Names.constant array -> Term.constr array -> int ->
Tacmach.tactic) : unit =
@@ -303,14 +303,14 @@ let generate_principle on_error
let funs_types = List.map (function ((_,_,_,types,_),_) -> types) fix_rec_l in
try
(* We then register the Inductive graphs of the functions *)
- Glob_term_to_relation.build_inductive names funs_args funs_types recdefs;
+ Glob_term_to_relation.build_inductive mp_dp 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 f_R_mut = Ident (Loc.ghost,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!")
@@ -326,11 +326,10 @@ let generate_principle on_error
in
let funs_kn = Array.of_list (List.map fname_kn fix_rec_l) in
let _ =
- list_map_i
+ 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
+ let princ = Indrec.lookup_eliminator (ind_kn,i) (InProp) in
+ let princ_type = Global.type_of_global_unsafe princ in
Functional_principles_types.generate_functional_principle
interactive_proof
princ_type
@@ -352,15 +351,11 @@ let generate_principle on_error
let register_struct is_rec (fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
match fixpoint_exprl with
| [((_,fname),_,bl,ret_type,body),_] when not is_rec ->
- let body = match body with | Some body -> body | None -> user_err_loc (dummy_loc,"Function",str "Body of Function must be given") in
- let ce,imps =
- Command.interp_definition bl None body (Some ret_type)
- in
- Command.declare_definition
- fname (Decl_kinds.Global,Decl_kinds.Definition)
- ce imps (fun _ _ -> ())
+ let body = match body with | Some body -> body | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given") in
+ Command.do_definition fname (Decl_kinds.Global,(*FIXME*)false,Decl_kinds.Definition)
+ bl None body (Some ret_type) (Lemmas.mk_hook (fun _ _ -> ()))
| _ ->
- Command.do_fixpoint fixpoint_exprl
+ Command.do_fixpoint Global false(*FIXME*) fixpoint_exprl
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
@@ -373,39 +368,39 @@ let generate_correction_proof_wf f_ref tcc_lemma_ref
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 type_of_f = Constrexpr_ops.prod_constr_expr ret_type args in
let rec_arg_num =
let names =
List.map
snd
- (Topconstr.names_of_local_assums args)
+ (Constrexpr_ops.names_of_local_assums args)
in
match wf_arg with
| None ->
- if List.length names = 1 then 1
+ if Int.equal (List.length names) 1 then 1
else error "Recursive argument must be specified"
| Some wf_arg ->
- list_index (Name wf_arg) names
+ List.index Name.equal (Name wf_arg) names
in
let unbounded_eq =
let f_app_args =
- Topconstr.CAppExpl
- (dummy_loc,
- (None,(Ident (dummy_loc,fname))) ,
+ Constrexpr.CAppExpl
+ (Loc.ghost,
+ (None,(Ident (Loc.ghost,fname)),None) ,
(List.map
(function
| _,Anonymous -> assert false
- | _,Name e -> (Topconstr.mkIdentC e)
+ | _,Name e -> (Constrexpr_ops.mkIdentC e)
)
- (Topconstr.names_of_local_assums args)
+ (Constrexpr_ops.names_of_local_assums args)
)
)
in
- Topconstr.CApp (dummy_loc,(None,Topconstr.mkRefC (Qualid (dummy_loc,(qualid_of_string "Logic.eq")))),
+ Constrexpr.CApp (Loc.ghost,(None,Constrexpr_ops.mkRefC (Qualid (Loc.ghost,(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
+ let eq = Constrexpr_ops.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
@@ -433,7 +428,7 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas
| None ->
begin
match args with
- | [Topconstr.LocalRawAssum ([(_,Name x)],k,t)] -> t,x
+ | [Constrexpr.LocalRawAssum ([(_,Name x)],k,t)] -> t,x
| _ -> error "Recursive argument must be specified"
end
| Some wf_args ->
@@ -441,15 +436,15 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas
match
List.find
(function
- | Topconstr.LocalRawAssum(l,k,t) ->
+ | Constrexpr.LocalRawAssum(l,k,t) ->
List.exists
- (function (_,Name id) -> id = wf_args | _ -> false)
+ (function (_,Name id) -> Id.equal id wf_args | _ -> false)
l
| _ -> false
)
args
with
- | Topconstr.LocalRawAssum(_,k,t) -> t,wf_args
+ | Constrexpr.LocalRawAssum(_,k,t) -> t,wf_args
| _ -> assert false
with Not_found -> assert false
in
@@ -457,31 +452,31 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas
match wf_rel_expr_opt with
| None ->
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")))
+ let make_dir l = DirPath.make (List.rev_map Id.of_string l) in
+ Libnames.Qualid (Loc.ghost,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)
+ Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC wf_arg]) in
+ Constrexpr_ops.mkLambdaC ([(Loc.ghost,Name wf_arg)],Constrexpr_ops.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])
+ Constrexpr_ops.mkAppC(Constrexpr_ops.mkRefC ltof,[wf_arg_type;fun_from_mes])
in
wf_rel_from_mes,true
| Some wf_rel_expr ->
let wf_rel_with_mes =
- let a = Names.id_of_string "___a" in
- let b = Names.id_of_string "___b" in
- Topconstr.mkLambdaC(
- [dummy_loc,Name a;dummy_loc,Name b],
- Topconstr.Default Lib.Explicit,
+ let a = Names.Id.of_string "___a" in
+ let b = Names.Id.of_string "___b" in
+ Constrexpr_ops.mkLambdaC(
+ [Loc.ghost,Name a;Loc.ghost,Name b],
+ Constrexpr.Default Explicit,
wf_arg_type,
- Topconstr.mkAppC(wf_rel_expr,
+ Constrexpr_ops.mkAppC(wf_rel_expr,
[
- Topconstr.mkAppC(wf_mes_expr,[Topconstr.mkIdentC a]);
- Topconstr.mkAppC(wf_mes_expr,[Topconstr.mkIdentC b])
+ Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC a]);
+ Constrexpr_ops.mkAppC(wf_mes_expr,[Constrexpr_ops.mkIdentC b])
])
)
in
@@ -493,124 +488,62 @@ let register_mes fname rec_impls wf_mes_expr wf_rel_expr_opt wf_arg using_lemmas
let map_option f = function
| None -> None
| Some v -> Some (f v)
-
-let decompose_lambda_n_assum_constr_expr =
- let rec decompose_lambda_n_assum_constr_expr acc n e =
- if n = 0 then (List.rev acc,e)
- else
- match e with
- | Topconstr.CLambdaN(_, [],e') -> decompose_lambda_n_assum_constr_expr acc n e'
- | Topconstr.CLambdaN(lambda_loc,(nal,bk,nal_type)::bl,e') ->
- let nal_length = List.length nal in
- if nal_length <= n
- then
- decompose_lambda_n_assum_constr_expr
- (Topconstr.LocalRawAssum(nal,bk,nal_type)::acc)
- (n - nal_length)
- (Topconstr.CLambdaN(lambda_loc,bl,e'))
- else
- let nal_keep,nal_expr = list_chop n nal in
- (List.rev (Topconstr.LocalRawAssum(nal_keep,bk,nal_type)::acc),
- Topconstr.CLambdaN(lambda_loc,(nal_expr,bk,nal_type)::bl,e')
- )
- | Topconstr.CLetIn(_, na,nav,e') ->
- decompose_lambda_n_assum_constr_expr
- (Topconstr.LocalRawDef(na,nav)::acc) (pred n) e'
- | _ -> error "Not enough product or assumption"
- in
- decompose_lambda_n_assum_constr_expr []
-
-let decompose_prod_n_assum_constr_expr =
- let rec decompose_prod_n_assum_constr_expr acc n e =
- (* Pp.msgnl (str "n := " ++ int n ++ fnl ()++ *)
- (* str "e := " ++ Ppconstr.pr_lconstr_expr e); *)
- if n = 0 then
- (* let _ = Pp.msgnl (str "return_type := " ++ Ppconstr.pr_lconstr_expr e) in *)
- (List.rev acc,e)
- else
- match e with
- | Topconstr.CProdN(_, [],e') -> decompose_prod_n_assum_constr_expr acc n e'
- | Topconstr.CProdN(lambda_loc,(nal,bk,nal_type)::bl,e') ->
- let nal_length = List.length nal in
- if nal_length <= n
- then
- (* let _ = Pp.msgnl (str "first case") in *)
- decompose_prod_n_assum_constr_expr
- (Topconstr.LocalRawAssum(nal,bk,nal_type)::acc)
- (n - nal_length)
- (if bl = [] then e' else (Topconstr.CLambdaN(lambda_loc,bl,e')))
- else
- (* let _ = Pp.msgnl (str "second case") in *)
- let nal_keep,nal_expr = list_chop n nal in
- (List.rev (Topconstr.LocalRawAssum(nal_keep,bk,nal_type)::acc),
- Topconstr.CLambdaN(lambda_loc,(nal_expr,bk,nal_type)::bl,e')
- )
- | Topconstr.CArrow(_,premisse,concl) ->
- (* let _ = Pp.msgnl (str "arrow case") in *)
- decompose_prod_n_assum_constr_expr
- (Topconstr.LocalRawAssum([dummy_loc,Names.Anonymous],
- Topconstr.Default Lib.Explicit,premisse)
- ::acc)
- (pred n)
- concl
- | Topconstr.CLetIn(_, na,nav,e') ->
- decompose_prod_n_assum_constr_expr
- (Topconstr.LocalRawDef(na,nav)::acc) (pred n) e'
- | _ -> error "Not enough product or assumption"
- in
- decompose_prod_n_assum_constr_expr []
+open Constrexpr
open Topconstr
-
-let id_of_name = function
- | Name id -> id
- | _ -> assert false
- let rec rebuild_bl (aux,assoc) bl typ =
+let make_assoc assoc l1 l2 =
+ let fold assoc a b = match a, b with
+ | (_, Name na), (_, Name id) -> Id.Map.add na id assoc
+ | _, _ -> assoc
+ in
+ List.fold_left2 fold assoc l1 l2
+
+let rec rebuild_bl (aux,assoc) bl typ =
match bl,typ with
| [], _ -> (List.rev aux,replace_vars_constr_expr assoc typ,assoc)
- | (Topconstr.LocalRawAssum(nal,bk,_))::bl',typ ->
+ | (Constrexpr.LocalRawAssum(nal,bk,_))::bl',typ ->
rebuild_nal (aux,assoc) bk bl' nal (List.length nal) typ
- | (Topconstr.LocalRawDef(na,_))::bl',CLetIn(_,_,nat,typ') ->
- rebuild_bl ((Topconstr.LocalRawDef(na,replace_vars_constr_expr assoc nat)::aux),assoc)
+ | (Constrexpr.LocalRawDef(na,_))::bl',Constrexpr.CLetIn(_,_,nat,typ') ->
+ rebuild_bl ((Constrexpr.LocalRawDef(na,replace_vars_constr_expr assoc nat)::aux),assoc)
bl' typ'
| _ -> assert false
and rebuild_nal (aux,assoc) bk bl' nal lnal typ =
match nal,typ with
| [], _ -> rebuild_bl (aux,assoc) bl' typ
- | na::nal,CArrow(_,nat,typ') ->
- rebuild_nal
- ((LocalRawAssum([na],bk,replace_vars_constr_expr assoc nat))::aux,assoc)
- bk bl' nal (pred lnal) typ'
| _,CProdN(_,[],typ) -> rebuild_nal (aux,assoc) bk bl' nal lnal typ
| _,CProdN(_,(nal',bk',nal't)::rest,typ') ->
let lnal' = List.length nal' in
if lnal' >= lnal
then
- let old_nal',new_nal' = list_chop lnal nal' in
- rebuild_bl ((LocalRawAssum(nal,bk,replace_vars_constr_expr assoc nal't)::aux),(List.rev_append (List.combine (List.map id_of_name (List.map snd old_nal')) (List.map id_of_name (List.map snd nal))) assoc)) bl'
- (if new_nal' = [] && rest = []
- then typ'
- else if new_nal' = []
- then CProdN(dummy_loc,rest,typ')
- else CProdN(dummy_loc,((new_nal',bk',nal't)::rest),typ'))
+ let old_nal',new_nal' = List.chop lnal nal' in
+ let nassoc = make_assoc assoc old_nal' nal in
+ let assum = LocalRawAssum(nal,bk,replace_vars_constr_expr assoc nal't) in
+ rebuild_bl ((assum :: aux), nassoc) bl'
+ (if List.is_empty new_nal' && List.is_empty rest
+ then typ'
+ else if List.is_empty new_nal'
+ then CProdN(Loc.ghost,rest,typ')
+ else CProdN(Loc.ghost,((new_nal',bk',nal't)::rest),typ'))
else
- let captured_nal,non_captured_nal = list_chop lnal' nal in
- rebuild_nal ((LocalRawAssum(captured_nal,bk,replace_vars_constr_expr assoc nal't)::aux), (List.rev_append (List.combine (List.map id_of_name (List.map snd captured_nal)) ((List.map id_of_name (List.map snd nal)))) assoc))
- bk bl' non_captured_nal (lnal - lnal') (CProdN(dummy_loc,rest,typ'))
+ let captured_nal,non_captured_nal = List.chop lnal' nal in
+ let nassoc = make_assoc assoc nal' captured_nal in
+ let assum = LocalRawAssum(captured_nal,bk,replace_vars_constr_expr assoc nal't) in
+ rebuild_nal ((assum :: aux), nassoc)
+ bk bl' non_captured_nal (lnal - lnal') (CProdN(Loc.ghost,rest,typ'))
| _ -> assert false
let rebuild_bl (aux,assoc) bl typ = rebuild_bl (aux,assoc) bl typ
let recompute_binder_list (fixpoint_exprl : (Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) =
let fixl,ntns = Command.extract_fixpoint_components false fixpoint_exprl in
- let ((_,_,typel),_) = Command.interp_fixpoint fixl ntns in
+ let ((_,_,typel),_,_) = Command.interp_fixpoint fixl ntns in
let constr_expr_typel =
- with_full_print (List.map (Constrextern.extern_constr false (Global.env ()))) typel in
+ with_full_print (List.map (Constrextern.extern_constr false (Global.env ()) Evd.empty)) typel in
let fixpoint_exprl_with_new_bl =
List.map2 (fun ((lna,(rec_arg_opt,rec_order),bl,ret_typ,opt_body),notation_list) fix_typ ->
- let new_bl',new_ret_type,_ = rebuild_bl ([],[]) bl fix_typ in
+ let new_bl',new_ret_type,_ = rebuild_bl ([],Id.Map.empty) bl fix_typ in
(((lna,(rec_arg_opt,rec_order),new_bl',new_ret_type,opt_body),notation_list):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list))
)
fixpoint_exprl constr_expr_typel
@@ -618,23 +551,24 @@ let recompute_binder_list (fixpoint_exprl : (Vernacexpr.fixpoint_expr * Vernacex
fixpoint_exprl_with_new_bl
-let do_generate_principle on_error register_built interactive_proof
+let do_generate_principle mp_dp on_error register_built interactive_proof
(fixpoint_exprl:(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list) list) :unit =
- List.iter (fun (_,l) -> if l <> [] then error "Function does not support notations for now") fixpoint_exprl;
+ List.iter (fun (_,l) -> if not (List.is_empty l) then error "Function does not support notations for now") fixpoint_exprl;
let _is_struct =
match fixpoint_exprl with
- | [((_,(wf_x,Topconstr.CWfRec wf_rel),_,_,_),_) as fixpoint_expr] ->
+ | [((_,(wf_x,Constrexpr.CWfRec wf_rel),_,_,_),_) as fixpoint_expr] ->
let ((((_,name),_,args,types,body)),_) as fixpoint_expr =
match recompute_binder_list [fixpoint_expr] with
| [e] -> e
| _ -> assert false
in
let fixpoint_exprl = [fixpoint_expr] in
- let body = match body with | Some body -> body | None -> user_err_loc (dummy_loc,"Function",str "Body of Function must be given") in
+ let body = match body with | Some body -> body | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given") in
let recdefs,rec_impls = build_newrecursive fixpoint_exprl in
let using_lemmas = [] in
let pre_hook =
generate_principle
+ mp_dp
on_error
true
register_built
@@ -645,7 +579,7 @@ let do_generate_principle on_error register_built interactive_proof
if register_built
then register_wf name rec_impls wf_rel (map_option snd wf_x) using_lemmas args types body pre_hook;
false
- |[((_,(wf_x,Topconstr.CMeasureRec(wf_mes,wf_rel_opt)),_,_,_),_) as fixpoint_expr] ->
+ |[((_,(wf_x,Constrexpr.CMeasureRec(wf_mes,wf_rel_opt)),_,_,_),_) as fixpoint_expr] ->
let ((((_,name),_,args,types,body)),_) as fixpoint_expr =
match recompute_binder_list [fixpoint_expr] with
| [e] -> e
@@ -654,9 +588,10 @@ let do_generate_principle on_error register_built interactive_proof
let fixpoint_exprl = [fixpoint_expr] in
let recdefs,rec_impls = build_newrecursive fixpoint_exprl in
let using_lemmas = [] in
- let body = match body with | Some body -> body | None -> user_err_loc (dummy_loc,"Function",str "Body of Function must be given") in
+ let body = match body with | Some body -> body | None -> user_err_loc (Loc.ghost,"Function",str "Body of Function must be given") in
let pre_hook =
generate_principle
+ mp_dp
on_error
true
register_built
@@ -670,7 +605,7 @@ let do_generate_principle on_error register_built interactive_proof
| _ ->
List.iter (function ((_na,(_,ord),_args,_body,_type),_not) ->
match ord with
- | Topconstr.CMeasureRec _ | Topconstr.CWfRec _ ->
+ | Constrexpr.CMeasureRec _ | Constrexpr.CWfRec _ ->
error
("Cannot use mutual definition with well-founded recursion or measure")
| _ -> ()
@@ -685,6 +620,7 @@ let do_generate_principle on_error register_built interactive_proof
let is_rec = List.exists (is_rec fix_names) recdefs in
if register_built then register_struct is_rec fixpoint_exprl;
generate_principle
+ mp_dp
on_error
false
register_built
@@ -697,18 +633,15 @@ let do_generate_principle on_error register_built interactive_proof
in
()
-open Topconstr
let rec add_args id new_args b =
match b with
- | CRef r ->
+ | CRef (r,_) ->
begin match r with
- | Libnames.Ident(loc,fname) when fname = id ->
- CAppExpl(dummy_loc,(None,r),new_args)
+ | Libnames.Ident(loc,fname) when Id.equal fname id ->
+ CAppExpl(Loc.ghost,(None,r,None),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)
+ | CFix _ | CCoFix _ -> anomaly ~label:"add_args " (Pp.str "todo")
| CProdN(loc,nal,b1) ->
CProdN(loc,
List.map (fun (nal,k,b2) -> (nal,k,add_args id new_args b2)) nal,
@@ -719,12 +652,12 @@ let rec add_args id new_args b =
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) ->
+ | CAppExpl(loc,(pf,r,us),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)
+ | Libnames.Ident(loc,fname) when Id.equal fname id ->
+ CAppExpl(loc,(pf,r,us),new_args@(List.map (add_args id new_args) exprl))
+ | _ -> CAppExpl(loc,(pf,r,us),List.map (add_args id new_args) exprl)
end
| CApp(loc,(pf,b),bl) ->
CApp(loc,(pf,add_args id new_args b),
@@ -733,7 +666,7 @@ let rec add_args id new_args b =
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,
+ (na, 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) ->
@@ -752,32 +685,29 @@ let rec add_args id new_args 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)
+ | CCast(loc,b1,b2) ->
+ CCast(loc,add_args id new_args b1,
+ Miscops.map_cast_type (add_args id new_args) b2)
| 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"
+ | CNotation _ -> anomaly ~label:"add_args " (Pp.str "CNotation")
+ | CGeneralization _ -> anomaly ~label:"add_args " (Pp.str "CGeneralization")
| CPrim _ -> b
- | CDelimiters _ -> anomaly "add_args : CDelimiters"
-exception Stop of Topconstr.constr_expr
+ | CDelimiters _ -> anomaly ~label:"add_args " (Pp.str "CDelimiters")
+exception Stop of Constrexpr.constr_expr
(* [chop_n_arrow n t] chops the [n] first arrows in [t]
- Acts on Topconstr.constr_expr
+ Acts on Constrexpr.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 :
+ | Constrexpr.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
@@ -796,8 +726,8 @@ let rec chop_n_arrow n t =
aux (n - nal_l) nal_ta'
else
let new_t' =
- Topconstr.CProdN(dummy_loc,
- ((snd (list_chop n nal)),k,t'')::nal_ta',t')
+ Constrexpr.CProdN(Loc.ghost,
+ ((snd (List.chop n nal)),k,t'')::nal_ta',t')
in
raise (Stop new_t')
in
@@ -806,13 +736,13 @@ let rec chop_n_arrow n t =
chop_n_arrow new_n t'
with Stop t -> t
end
- | _ -> anomaly "Not enough products"
+ | _ -> anomaly (Pp.str "Not enough products")
-let rec get_args b t : Topconstr.local_binder list *
- Topconstr.constr_expr * Topconstr.constr_expr =
+let rec get_args b t : Constrexpr.local_binder list *
+ Constrexpr.constr_expr * Constrexpr.constr_expr =
match b with
- | Topconstr.CLambdaN (loc, (nal_ta), b') ->
+ | Constrexpr.CLambdaN (loc, (nal_ta), b') ->
begin
let n =
(List.fold_left (fun n (nal,_,_) ->
@@ -820,7 +750,7 @@ let rec get_args b t : Topconstr.local_binder list *
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''
+ (Constrexpr.LocalRawAssum (nal,k,ta))) nal_ta)@nal_tas, b'',t''
end
| _ -> [],b,t
@@ -836,17 +766,15 @@ let make_graph (f_ref:global_reference) =
| _ -> raise (UserError ("", str "Not a function reference") )
in
Dumpglob.pause ();
- (match body_of_constant c_body with
+ (match Global.body_of_constant_body c_body with
| None -> error "Cannot build a graph over an axiom !"
- | Some b ->
+ | Some body ->
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)
+ with_full_print (fun () ->
+ (Constrextern.extern_constr false env Evd.empty body,
+ Constrextern.extern_type false env Evd.empty
+ ((*FIXNE*) Typeops.type_of_constant_type env c_body.const_type)
)
)
()
@@ -854,7 +782,7 @@ let make_graph (f_ref:global_reference) =
let (nal_tas,b,t) = get_args extern_body extern_type in
let expr_list =
match b with
- | Topconstr.CFix(loc,l_id,fixexprl) ->
+ | Constrexpr.CFix(loc,l_id,fixexprl) ->
let l =
List.map
(fun (id,(n,recexp),bl,t,b) ->
@@ -863,34 +791,34 @@ let make_graph (f_ref:global_reference) =
List.flatten
(List.map
(function
- | Topconstr.LocalRawDef (na,_)-> []
- | Topconstr.LocalRawAssum (nal,_,_) ->
+ | Constrexpr.LocalRawDef (na,_)-> []
+ | Constrexpr.LocalRawAssum (nal,_,_) ->
List.map
(fun (loc,n) ->
- CRef(Libnames.Ident(loc, Nameops.out_name n)))
+ CRef(Libnames.Ident(loc, Nameops.out_name n),None))
nal
)
nal_tas
)
in
let b' = add_args (snd id) new_args b in
- (((id, ( Some (dummy_loc,rec_id),CStructRec),nal_tas@bl,t,Some b'),[]):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list))
+ (((id, ( Some (Loc.ghost,rec_id),CStructRec),nal_tas@bl,t,Some b'),[]):(Vernacexpr.fixpoint_expr * Vernacexpr.decl_notation list))
)
fixexprl
in
l
| _ ->
- let id = id_of_label (con_label c) in
- [((dummy_loc,id),(None,Topconstr.CStructRec),nal_tas,t,Some b),[]]
+ let id = Label.to_id (con_label c) in
+ [((Loc.ghost,id),(None,Constrexpr.CStructRec),nal_tas,t,Some b),[]]
in
- do_generate_principle error_error false false expr_list;
- (* We register the infos *)
let mp,dp,_ = repr_con c in
+ do_generate_principle (Some (mp,dp)) error_error false false expr_list;
+ (* We register the infos *)
List.iter
- (fun (((_,id),_,_,_,_),_) -> add_Function false (make_con mp dp (label_of_id id)))
+ (fun (((_,id),_,_,_,_),_) -> add_Function false (make_con mp dp (Label.of_id id)))
expr_list);
Dumpglob.continue ()
-let do_generate_principle = do_generate_principle warning_error true
+let do_generate_principle = do_generate_principle None warning_error true
diff --git a/plugins/funind/indfun.mli b/plugins/funind/indfun.mli
index e65b5808..e7206914 100644
--- a/plugins/funind/indfun.mli
+++ b/plugins/funind/indfun.mli
@@ -1,11 +1,4 @@
-open Util
-open Names
-open Term
-open Pp
-open Indfun_common
-open Libnames
-open Glob_term
-open Declarations
+open Misctypes
val do_generate_principle :
bool ->
@@ -16,9 +9,9 @@ val do_generate_principle :
val functional_induction :
bool ->
Term.constr ->
- (Term.constr * Term.constr Glob_term.bindings) option ->
- Genarg.intro_pattern_expr Util.located option ->
+ (Term.constr * Term.constr bindings) option ->
+ Tacexpr.or_and_intro_pattern option ->
Proof_type.goal Tacmach.sigma -> Proof_type.goal list Evd.sigma
-val make_graph : Libnames.global_reference -> unit
+val make_graph : Globnames.global_reference -> unit
diff --git a/plugins/funind/indfun_common.ml b/plugins/funind/indfun_common.ml
index 827191b1..76f8c6d2 100644
--- a/plugins/funind/indfun_common.ml
+++ b/plugins/funind/indfun_common.ml
@@ -1,9 +1,9 @@
open Names
open Pp
-
open Libnames
-
-let mk_prefix pre id = id_of_string (pre^(string_of_id id))
+open Globnames
+open Refiner
+let mk_prefix pre id = Id.of_string (pre^(Id.to_string 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"
@@ -12,10 +12,7 @@ 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_id avoid s = Namegen.next_ident_away_in_goal (Id.of_string s) avoid
let fresh_name avoid s = Name (fresh_id avoid s)
@@ -29,7 +26,7 @@ let array_get_start a =
(Array.length a - 1)
(fun i -> a.(i))
with Invalid_argument "index out of bounds" ->
- invalid_argument "array_get_start"
+ invalid_arg "array_get_start"
let id_of_name = function
Name id -> id
@@ -51,10 +48,8 @@ let locate_constant ref =
let locate_with_msg msg f x =
- try
- f x
- with
- | Not_found -> raise (Util.UserError("", msg))
+ try f x
+ with Not_found -> raise (Errors.UserError("", msg))
let filter_map filter f =
@@ -78,7 +73,7 @@ let chop_rlambda_n =
| Glob_term.GLambda(_,name,k,t,b) -> chop_lambda_n ((name,t,false)::acc) (n-1) b
| Glob_term.GLetIn(_,name,v,b) -> chop_lambda_n ((name,v,true)::acc) (n-1) b
| _ ->
- raise (Util.UserError("chop_rlambda_n",
+ raise (Errors.UserError("chop_rlambda_n",
str "chop_rlambda_n: Not enough Lambdas"))
in
chop_lambda_n []
@@ -90,7 +85,7 @@ let chop_rprod_n =
else
match rt with
| Glob_term.GProd(_,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"))
+ | _ -> raise (Errors.UserError("chop_rprod_n",str "chop_rprod_n: Not enough products"))
in
chop_prod_n []
@@ -111,34 +106,27 @@ let list_add_set_eq eq_fun x l =
let const_of_id id =
let _,princ_ref =
- qualid_of_reference (Libnames.Ident (Util.dummy_loc,id))
+ qualid_of_reference (Libnames.Ident (Loc.ghost,id))
in
try Nametab.locate_constant princ_ref
- with Not_found -> Util.error ("cannot find "^ string_of_id id)
+ with Not_found -> Errors.error ("cannot find "^ Id.to_string id)
let def_of_const t =
match (Term.kind_of_term t) with
Term.Const sp ->
- (try (match Declarations.body_of_constant (Global.lookup_constant sp) with
- | Some c -> Declarations.force c
+ (try (match Environ.constant_opt_value_in (Global.env()) sp with
+ | Some c -> c
| _ -> assert false)
- with e when Errors.noncritical e -> assert false)
+ with Not_found -> 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 dp = Names.DirPath.make (List.rev_map Id.of_string sl) in
+ Nametab.locate (make_qualid dp (Id.of_string s))
let eq = lazy(coq_constant "eq")
let refl_equal = lazy(coq_constant "eq_refl")
@@ -147,47 +135,40 @@ 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 definition_message = Declare.definition_message
-let save with_clean id const (locality,kind) hook =
- let {const_entry_body = pft;
- const_entry_secctx = _;
- const_entry_type = tpo;
- const_entry_opaque = opacity } = const in
+let get_locality = function
+| Discharge -> true
+| Local -> true
+| Global -> false
+
+let save with_clean id const (locality,_,kind) hook =
+ let fix_exn = Future.fix_exn_of const.Entries.const_entry_body 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
+ | Discharge when Lib.sections_are_opened () ->
+ let k = Kindops.logical_kind_of_goal_kind kind in
+ let c = SectionLocalDef const 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
+ | Discharge | Local | Global ->
+ let local = get_locality locality in
+ let k = Kindops.logical_kind_of_goal_kind kind in
+ let kn = declare_constant id ~local (DefinitionEntry const, k) in
+ (locality, ConstRef kn)
+ in
if with_clean then Pfedit.delete_current_proof ();
- hook l r;
+ Ephemeron.iter_opt hook (fun f -> Lemmas.call_hook fix_exn f l r);
definition_message id
let cook_proof _ =
- let (id,(entry,_,strength,hook)) = Pfedit.cook_proof (fun _ -> ()) in
- (id,(entry,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 (id,(entry,_,strength)) = Pfedit.cook_proof () in
+ (id,(entry,strength))
let get_proof_clean do_reduce =
let result = cook_proof do_reduce in
@@ -197,7 +178,8 @@ let get_proof_clean do_reduce =
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
+ 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;
@@ -248,8 +230,9 @@ type function_info =
(* let function_table = ref ([] : function_db) *)
-let from_function = ref Cmap.empty
-let from_graph = ref Indmap.empty
+let from_function = Summary.ref Cmap_env.empty ~name:"functions_db_fn"
+let from_graph = Summary.ref Indmap.empty ~name:"functions_db_gr"
+
(*
let rec do_cache_info finfo = function
| [] -> raise Not_found
@@ -272,15 +255,14 @@ let cache_Function (_,(finfos)) =
*)
let cache_Function (_,finfos) =
- from_function := Cmap.add finfos.function_constant finfos !from_function;
+ from_function := Cmap_env.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)
+ let do_subst_con c = Mod_subst.subst_constant subst c
+ and do_subst_ind i = Mod_subst.subst_ind subst i
in
let function_constant' = do_subst_con finfos.function_constant in
let graph_ind' = do_subst_ind finfos.graph_ind in
@@ -346,22 +328,29 @@ let discharge_Function (_,finfos) =
}
open Term
+
+let pr_ocst c =
+ Option.fold_right (fun v acc -> Printer.pr_lconstr (mkConst v)) c (mt ())
+
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 e when Errors.noncritical e -> 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 ()
+ str "function_constant := " ++
+ Printer.pr_lconstr (mkConst f_info.function_constant)++ fnl () ++
+ str "function_constant_type := " ++
+ (try
+ Printer.pr_lconstr
+ (Global.type_of_global_unsafe (ConstRef f_info.function_constant))
+ with e when Errors.noncritical e -> mt ()) ++ fnl () ++
+ str "equation_lemma := " ++ pr_ocst f_info.equation_lemma ++ fnl () ++
+ str "completeness_lemma :=" ++ pr_ocst f_info.completeness_lemma ++ fnl () ++
+ str "correctness_lemma := " ++ pr_ocst f_info.correctness_lemma ++ fnl () ++
+ str "rect_lemma := " ++ pr_ocst f_info.rect_lemma ++ fnl () ++
+ str "rec_lemma := " ++ pr_ocst f_info.rec_lemma ++ fnl () ++
+ str "prop_lemma := " ++ pr_ocst f_info.prop_lemma ++ 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 l = Cmap_env.fold (fun k v acc -> v::acc) tb [] in
+ Pp.prlist_with_sep fnl pr_info l
let in_Function : function_info -> Libobject.obj =
Libobject.declare_object
@@ -375,36 +364,16 @@ let in_Function : function_info -> Libobject.obj =
}
-
-(* 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"
+ (match Nametab.locate (qualid_of_ident id) with ConstRef c -> c | _ -> Errors.anomaly (Pp.str "Not a constant")
)
with Not_found -> None
let find_Function_infos f =
- Cmap.find f !from_function
+ Cmap_env.find f !from_function
let find_Function_of_graph ind =
@@ -416,7 +385,7 @@ let update_Function finfo =
let add_Function is_general f =
- let f_id = id_of_label (con_label f) in
+ let f_id = Label.to_id (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)
@@ -425,7 +394,7 @@ let add_Function is_general f =
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"
+ with | IndRef ind -> ind | _ -> Errors.anomaly (Pp.str "Not an inductive")
in
let finfos =
{ function_constant = f;
@@ -475,8 +444,7 @@ let function_debug_sig =
let _ = declare_bool_option function_debug_sig
-let do_observe () =
- !function_debug = true
+let do_observe () = !function_debug
@@ -499,25 +467,37 @@ 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 when Errors.noncritical e -> raise (ToShow e)
-
let jmeq () =
try
- (Coqlib.check_required_library ["Coq";"Logic";"JMeq"];
- init_constant ["Logic";"JMeq"] "JMeq")
- with e when Errors.noncritical e -> raise (ToShow e)
-
-let jmeq_rec () =
- try
- Coqlib.check_required_library ["Coq";"Logic";"JMeq"];
- init_constant ["Logic";"JMeq"] "JMeq_rec"
+ Coqlib.check_required_library Coqlib.jmeq_module_name;
+ Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq"
with e when Errors.noncritical e -> raise (ToShow e)
let jmeq_refl () =
try
- Coqlib.check_required_library ["Coq";"Logic";"JMeq"];
- init_constant ["Logic";"JMeq"] "JMeq_refl"
+ Coqlib.check_required_library Coqlib.jmeq_module_name;
+ Coqlib.gen_constant "Function" ["Logic";"JMeq"] "JMeq_refl"
with e when Errors.noncritical e -> raise (ToShow e)
+
+let h_intros l =
+ tclMAP (fun x -> Proofview.V82.of_tactic (Tactics.Simple.intro x)) l
+
+let h_id = Id.of_string "h"
+let hrec_id = Id.of_string "hrec"
+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 ltof_ref = function () -> (find_reference ["Coq";"Arith";"Wf_nat"] "ltof")
+
+let evaluable_of_global_reference r = (* Tacred.evaluable_of_global_reference (Global.env ()) *)
+ match r with
+ ConstRef sp -> EvalConstRef sp
+ | VarRef id -> EvalVarRef id
+ | _ -> assert false;;
+
+let list_rewrite (rev:bool) (eqs: (constr*bool) list) =
+ tclREPEAT
+ (List.fold_right
+ (fun (eq,b) i -> tclORELSE (Proofview.V82.of_tactic ((if b then Equality.rewriteLR else Equality.rewriteRL) eq)) i)
+ (if rev then (List.rev eqs) else eqs) (tclFAIL 0 (mt())));;
diff --git a/plugins/funind/indfun_common.mli b/plugins/funind/indfun_common.mli
index e0076735..67ddf374 100644
--- a/plugins/funind/indfun_common.mli
+++ b/plugins/funind/indfun_common.mli
@@ -5,23 +5,21 @@ 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 mk_rel_id : Id.t -> Id.t
+val mk_correct_id : Id.t -> Id.t
+val mk_complete_id : Id.t -> Id.t
+val mk_equation_id : Id.t -> Id.t
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 fresh_id : Id.t list -> string -> Id.t
+val fresh_name : Id.t list -> string -> Name.t
+val get_name : Id.t list -> ?default:string -> Name.t -> Name.t
val array_get_start : 'a array -> 'a array
-val id_of_name : name -> identifier
+val id_of_name : Name.t -> Id.t
val locate_ind : Libnames.reference -> inductive
val locate_constant : Libnames.reference -> constant
@@ -36,38 +34,31 @@ val list_add_set_eq :
('a -> 'a -> bool) -> 'a -> 'a list -> 'a list
val chop_rlambda_n : int -> Glob_term.glob_constr ->
- (name*Glob_term.glob_constr*bool) list * Glob_term.glob_constr
+ (Name.t*Glob_term.glob_constr*bool) list * Glob_term.glob_constr
val chop_rprod_n : int -> Glob_term.glob_constr ->
- (name*Glob_term.glob_constr) list * Glob_term.glob_constr
+ (Name.t*Glob_term.glob_constr) list * Glob_term.glob_constr
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 const_of_id: Id.t -> 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]
-*)
-
-val new_save_named : bool -> unit
-
-val save : bool -> identifier -> Entries.definition_entry -> Decl_kinds.goal_kind ->
- Tacexpr.declaration_hook -> unit
+val save : bool -> Id.t -> Entries.definition_entry -> Decl_kinds.goal_kind ->
+ unit Lemmas.declaration_hook Ephemeron.key -> 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)
+ Names.Id.t *
+ (Entries.definition_entry * Decl_kinds.goal_kind)
-(* [with_full_print f a] applies [f] to [a] in full printing environment
+(* [with_full_print f a] applies [f] to [a] in full printing environment.
This function preserves the print settings
*)
@@ -112,3 +103,14 @@ exception Defining_principle of exn
exception ToShow of exn
val is_strict_tcc : unit -> bool
+
+val h_intros: Names.Id.t list -> Proof_type.tactic
+val h_id : Names.Id.t
+val hrec_id : Names.Id.t
+val acc_inv_id : Term.constr Util.delayed
+val ltof_ref : Globnames.global_reference Util.delayed
+val well_founded_ltof : Term.constr Util.delayed
+val acc_rel : Term.constr Util.delayed
+val well_founded : Term.constr Util.delayed
+val evaluable_of_global_reference : Globnames.global_reference -> Names.evaluable_global_reference
+val list_rewrite : bool -> (Term.constr*bool) list -> Proof_type.tactic
diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml
index c770c7ce..0c7b0a0b 100644
--- a/plugins/funind/invfun.ml
+++ b/plugins/funind/invfun.ml
@@ -1,39 +1,40 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
open Tacexpr
open Declarations
+open Errors
open Util
open Names
open Term
+open Vars
open Pp
-open Libnames
+open Globnames
open Tacticals
open Tactics
open Indfun_common
open Tacmach
-open Sign
-open Hiddentac
+open Misctypes
(* Some pretty printing function for debugging purpose *)
let pr_binding prc =
function
- | loc, Glob_term.NamedHyp id, c -> hov 1 (Ppconstr.pr_id id ++ str " := " ++ Pp.cut () ++ prc c)
- | loc, Glob_term.AnonHyp n, c -> hov 1 (int n ++ str " := " ++ Pp.cut () ++ prc c)
+ | 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
- | Glob_term.ImplicitBindings l ->
+ | ImplicitBindings l ->
brk (1,1) ++ str "with" ++ brk (1,1) ++
- Util.prlist_with_sep spc prc l
- | Glob_term.ExplicitBindings l ->
+ pr_sequence prc l
+ | ExplicitBindings l ->
brk (1,1) ++ str "with" ++ brk (1,1) ++
- Util.prlist_with_sep spc (fun b -> str"(" ++ pr_binding prlc b ++ str")") l
- | Glob_term.NoBindings -> mt ()
+ pr_sequence (fun b -> str"(" ++ pr_binding prlc b ++ str")") l
+ | NoBindings -> mt ()
let pr_with_bindings prc prlc (c,bl) =
@@ -45,17 +46,17 @@ 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 msgnl = Pp.msgnl *)
let observe strm =
if do_observe ()
- then Pp.msgnl strm
+ then Pp.msg_debug strm
else ()
-let observennl strm =
+(*let observennl strm =
if do_observe ()
then begin Pp.msg strm;Pp.pp_flush () end
- else ()
+ else ()*)
let do_observe_tac s tac g =
@@ -64,22 +65,25 @@ let do_observe_tac s tac g =
with e when Errors.noncritical e -> assert false
in
try
- let v = tac g in msgnl (goal ++ fnl () ++ s ++(str " ")++(str "finished")); v
+ let v = tac g in
+ msgnl (goal ++ fnl () ++ s ++(str " ")++(str "finished")); v
with reraise ->
- let e' = Cerrors.process_vernac_interp_error reraise in
+ let reraise = Errors.push reraise in
+ let e = Cerrors.process_vernac_interp_error reraise in
msgnl (str "observation "++ s++str " raised exception " ++
- Errors.print e' ++ str " on goal " ++ goal );
- raise reraise;;
-
+ Errors.iprint e ++ str " on goal " ++ goal );
+ iraise reraise;;
-let observe_tac_msg s tac g =
- if do_observe ()
+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 =
- observe_tac_msg (str s) tac g
+
+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 =
@@ -109,57 +113,47 @@ let id_to_constr id =
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 gr,u = destInd graph in
+ let graph_arity = Inductive.type_of_inductive (Global.env())
+ (Global.lookup_inductive gr, u) in
let ctxt,_ = decompose_prod_assum graph_arity in
let fun_ctxt,res_type =
match ctxt with
- | [] | [_] -> anomaly "Not a valid context"
+ | [] | [_] -> anomaly (Pp.str "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)
+ let rec args_from_decl i accu = function
+ | [] -> accu
+ | (_, Some _, _) :: l ->
+ args_from_decl (succ i) accu l
+ | _ :: l ->
+ let t = mkRel i in
+ args_from_decl (succ i) (t :: accu) l
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
+ let filter = function (Name id,_,_) -> Some id | (Anonymous,_,_) -> None in
+ let named_ctxt = 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
(*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
+ let args_as_rels = Array.of_list (args_from_decl 1 [] fun_ctxt) 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 make_eq () =
+(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq ())
+ in
let res_eq_f_of_args =
- mkApp(Coqlib.build_coq_eq (),[|lift 2 res_type;mkRel 1;mkRel 2|])
+ mkApp(make_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
+ let args_and_res_as_rels = Array.of_list (args_from_decl 3 [] fun_ctxt) in
+ let args_and_res_as_rels = Array.append args_and_res_as_rels [|mkRel 1|] in
+ let graph_applied = 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*)
@@ -178,7 +172,7 @@ let generate_type g_to_f f graph i =
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
+ let f_as_constant,u = match kind_of_term f with
| Const c' -> c'
| _ -> error "Must be used with a function"
in
@@ -195,7 +189,7 @@ let find_induction_principle f =
(* let fname = *)
(* match kind_of_term f with *)
(* | Const c' -> *)
-(* id_of_label (con_label c') *)
+(* Label.to_id (con_label c') *)
(* | _ -> error "Must be used with a function" *)
(* in *)
@@ -217,6 +211,11 @@ let rec generate_fresh_id x avoid i =
let id = Namegen.next_ident_away_in_goal x avoid in
id::(generate_fresh_id x (id::avoid) (pred i))
+let make_eq () =
+(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq ())
+let make_eq_refl () =
+(*FIXME*) Universes.constr_of_global (Coqlib.build_coq_eq_refl ())
+
(* [prove_fun_correct functional_induction funs_constr graphs_constr schemes lemmas_types_infos i ]
is the tactic used to prove correctness lemma.
@@ -248,11 +247,266 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
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 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 -> Loc.ghost, IntroNaming (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 eq_ind = make_eq () in
+ let eq_construct = mkConstructUi (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 -> Id.Set.add id acc
+ | _ -> anomaly (Pp.str "Not an identifier")
+ )
+ (List.nth intro_pats (pred i))
+ Id.Set.empty
+ in
+ let pre_args g =
+ List.fold_right
+ (fun (id,b,t) pre_args ->
+ if Id.Set.mem id this_branche_ids
+ then
+ match b with
+ | None -> id::pre_args
+ | Some b -> pre_args
+ else pre_args
+ )
+ (pf_hyps g)
+ ([])
+ in
+ let pre_args g = List.rev (pre_args g) in
+ let pre_tac g =
+ List.fold_right
+ (fun (id,b,t) pre_tac ->
+ if Id.Set.mem id this_branche_ids
+ then
+ match b with
+ | None -> pre_tac
+ | Some b ->
+ tclTHEN (h_reduce (Glob_term.Unfold([Glob_term.AllOccurrences,EvalVarRef id])) allHyps) pre_tac
+ else pre_tac
+ )
+ (pf_hyps g)
+ tclIDTAC
+ in
+*)
+ let pre_args =
+ List.fold_right
+ (fun (_,pat) acc ->
+ match pat with
+ | IntroNaming (IntroIdentifier id) -> id::acc
+ | _ -> anomaly (Pp.str "Not an identifier")
+ )
+ (List.nth intro_pats (pred i))
+ []
+ in
+ (* and get the real args of the branch by unfolding the defined constant *)
+ (*
+ 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 g =
+ 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 ->
+ (args.(2)::(mkApp(mkVar hid,[|args.(2);(mkApp(eq_construct,[|args.(0);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 g =
+ let params_id = fst (List.chop princ_infos.nparams args_names) in
+ (List.map mkVar params_id)@((constructor_args g))
+ 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 g = applist((mkConstruct(constructor)),constructor_args g) 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
+ [
+ observe_tac("h_intro_patterns ") (let l = (List.nth intro_pats (pred i)) in
+ match l with
+ | [] -> tclIDTAC
+ | _ -> Proofview.V82.of_tactic (intro_patterns l));
+ (* unfolding of all the defined variables introduced by this branch *)
+ (* observe_tac "unfolding" pre_tac; *)
+ (* $zeta$ normalizing of the conclusion *)
+ reduce
+ (Genredexpr.Cbv
+ { Redops.all_flags with
+ Genredexpr.rDelta = false ;
+ Genredexpr.rConst = []
+ }
+ )
+ Locusops.onConcl;
+ observe_tac ("toto ") tclIDTAC;
+
+ (* introducing the 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)));
+ (* Conclusion *)
+ observe_tac "exact" (fun g -> Proofview.V82.of_tactic (exact_check (app_constructor g)) g)
+ ]
+ )
+ g
+ in
+ (* end of branche proof *)
+ let lemmas =
+ Array.map
+ (fun (_,(ctxt,concl)) ->
+ match ctxt with
+ | [] | [_] | [_;_] -> anomaly (Pp.str "bad context")
+ | hres::res::(x,_,t)::ctxt ->
+ Termops.it_mkLambda_or_LetIn
+ (Termops.it_mkProd_or_LetIn concl [hres;res])
+ ((x,None,t)::ctxt)
+ )
+ lemmas_types_infos
+ in
+ 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
+ (*(Loc.ghost,Glob_term.NamedHyp id,p)*)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
+ (*(Loc.ghost,Glob_term.NamedHyp id,(nf_zeta p))*) (nf_zeta p)::bindings,id::avoid)
+ ([],avoid)
+ princ_infos.predicates
+ (lemmas)))
+ in
+ (* Glob_term.ExplicitBindings *) (params_bindings@lemmas_bindings)
+ in
+ tclTHENSEQ
+ [
+ observe_tac "principle" (Proofview.V82.of_tactic (assert_by
+ (Name principle_id)
+ princ_type
+ (exact_check f_principle)));
+ observe_tac "intro args_names" (tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) args_names);
+ (* observe_tac "titi" (pose_proof (Name (Id.of_string "__")) (Reductionops.nf_beta Evd.empty ((mkApp (mkVar principle_id,Array.of_list bindings))))); *)
+ observe_tac "idtac" tclIDTAC;
+ tclTHEN_i
+ (observe_tac "functional_induction" (
+ (fun gl ->
+ let term = mkApp (mkVar principle_id,Array.of_list bindings) in
+ let gl', _ty = pf_eapply Typing.e_type_of gl term in
+ Proofview.V82.of_tactic (apply term) gl')
+ ))
+ (fun i g -> observe_tac ("proving branche "^string_of_int i) (prove_branche i) g )
+ ]
+ g
+
+
+(*
+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"
+ | [] | [_] | [_;_] -> anomaly (Pp.str "bad context")
| hres::res::(x,_,t)::ctxt ->
Termops.it_mkLambda_or_LetIn
(Termops.it_mkProd_or_LetIn concl [hres;res])
@@ -270,13 +524,13 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
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 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 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
@@ -285,44 +539,43 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
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))))
+ (fun id -> Loc.ghost, 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 (Loc.ghost,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 this_branche_ids empty add i =
- List.fold_right
- (fun (_,pat) acc ->
- match pat with
- | Genarg.IntroIdentifier id -> add id acc
- | _ -> anomaly "Not an identifier"
- )
- (List.nth intro_pats (pred i))
- empty
- in
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 -> Id.Set.add id acc
+ | _ -> anomaly (Pp.str "Not an identifier")
+ )
+ (List.nth intro_pats (pred i))
+ Id.Set.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 Idset.empty Idset.add i)
+ if Id.Set.mem id this_branche_ids
then
match b with
- | None ->
- (id::pre_args,pre_tac)
+ | None -> (id::pre_args,pre_tac)
| Some b ->
(pre_args,
- tclTHEN (h_reduce (Glob_term.Unfold([Glob_term.all_occurrences_expr,EvalVarRef id])) allHyps) pre_tac
+ tclTHEN (h_reduce (Glob_term.Unfold([Glob_term.AllOccurrences,EvalVarRef id])) allHyps) pre_tac
)
-
else (pre_args,pre_tac)
)
(pf_hyps g)
@@ -333,7 +586,6 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
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 =
@@ -350,9 +602,9 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
| 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)
+ 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
@@ -362,7 +614,7 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
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
+ 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
@@ -390,11 +642,11 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
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
+ match generate_fresh_id (Id.of_string "z") (ids(* @this_branche_ids *)) 2 with
| [res;hres] -> res,hres
| _ -> assert false
in
- observe_tac_msg (str "constructor := " ++ Printer.pr_lconstr_env (pf_env g) app_constructor)
+ observe (str "constructor := " ++ Printer.pr_lconstr_env (pf_env g) app_constructor);
(
tclTHENSEQ
[
@@ -414,13 +666,13 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
(* replacing [res] with its value *)
observe_tac "rewriting res value" (Equality.rewriteLR (mkVar hres));
(* Conclusion *)
- observe_tac "exact" (h_exact app_constructor)
+ observe_tac "exact" (exact_check app_constructor)
]
)
g
in
(* end of branche proof *)
- let param_names = fst (list_chop princ_infos.nparams args_names) in
+ 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
@@ -431,7 +683,7 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
List.fold_left2
(fun (bindings,avoid) (x,_,_) p ->
let id = Namegen.next_ident_away (Nameops.out_name x) avoid in
- (dummy_loc,Glob_term.NamedHyp id,p)::bindings,id::avoid
+ (Loc.ghost,Glob_term.NamedHyp id,p)::bindings,id::avoid
)
([],pf_ids_of_hyps g)
princ_infos.params
@@ -441,7 +693,7 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
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,Glob_term.NamedHyp id,(nf_zeta p))::bindings,id::avoid)
+ (Loc.ghost,Glob_term.NamedHyp id,(nf_zeta p))::bindings,id::avoid)
([],avoid)
princ_infos.predicates
(lemmas)))
@@ -453,18 +705,21 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
observe_tac "principle" (assert_by
(Name principle_id)
princ_type
- (h_exact f_principle));
+ (exact_check 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));
- h_apply false false [dummy_loc,(mkVar principle_id,bindings)] g
+ 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)
- (tclTHEN (tclMAP h_intro (this_branche_ids [] (fun a b -> a::b) i)) (prove_branche i)) 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]
@@ -472,8 +727,8 @@ let prove_fun_correct functional_induction funs_constr graphs_constr schemes lem
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])
+ | (id,None,t) when not (Id.equal id hyp) &&
+ (Termops.occur_var (pf_env g) x t) -> tclTHEN (Tactics.Simple.generalize [mkVar id]) (thin [id])
| _ -> tclIDTAC
)
(pf_hyps g)
@@ -490,7 +745,7 @@ 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
+ let eq_ind = make_eq () in
match kind_of_term (pf_concl g) with
| Prod(_,t,t') ->
begin
@@ -498,66 +753,79 @@ and intros_with_rewrite_aux : tactic =
| 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
-
+ 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[
+ unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))];
+ tclMAP (fun id -> tclTRY(unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(1)))] ((destVar 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[
+ unfold_in_concl [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))];
+ tclMAP (fun id -> tclTRY(unfold_in_hyp [(Locus.AllOccurrences, Names.EvalVarRef (destVar args.(2)))] ((destVar args.(2)),Locus.InHyp) ))
+ (pf_ids_of_hyps g);
+ 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;
+ 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;
- tclTRY (Equality.rewriteLR (mkVar id));
+ tclTRY (Proofview.V82.of_tactic (Equality.rewriteLR (mkVar id)));
intros_with_rewrite
]
g
else if isVar args.(2)
then
- let id = pf_get_new_id (id_of_string "y") g in
- tclTHENSEQ [ h_intro id;
+ 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;
- tclTRY (Equality.rewriteRL (mkVar id));
+ tclTRY (Proofview.V82.of_tactic (Equality.rewriteRL (mkVar id)));
intros_with_rewrite
]
g
else
begin
- let id = pf_get_new_id (id_of_string "y") g in
+ let id = pf_get_new_id (Id.of_string "y") g in
tclTHENSEQ[
- h_intro id;
- tclTRY (Equality.rewriteLR (mkVar id));
+ 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 ()) ->
- Tauto.tauto g
+ Proofview.V82.of_tactic Tauto.tauto g
| Case(_,_,v,_) ->
tclTHENSEQ[
- h_case false (v,Glob_term.NoBindings);
+ Proofview.V82.of_tactic (simplest_case v);
intros_with_rewrite
] g
| LetIn _ ->
tclTHENSEQ[
- h_reduce
- (Glob_term.Cbv
- {Glob_term.all_flags
- with Glob_term.rDelta = false;
+ reduce
+ (Genredexpr.Cbv
+ {Redops.all_flags
+ with Genredexpr.rDelta = false;
})
- onConcl
+ Locusops.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
+ let id = pf_get_new_id (Id.of_string "y") g in
+ tclTHENSEQ [ Proofview.V82.of_tactic (Simple.intro id);intros_with_rewrite] g
end
| LetIn _ ->
tclTHENSEQ[
- h_reduce
- (Glob_term.Cbv
- {Glob_term.all_flags
- with Glob_term.rDelta = false;
+ reduce
+ (Genredexpr.Cbv
+ {Redops.all_flags
+ with Genredexpr.rDelta = false;
})
- onConcl
+ Locusops.onConcl
;
intros_with_rewrite
] g
@@ -569,14 +837,14 @@ let rec reflexivity_with_destruct_cases g =
match kind_of_term (snd (destApp (pf_concl g))).(2) with
| Case(_,_,v,_) ->
tclTHENSEQ[
- h_case false (v,Glob_term.NoBindings);
- intros;
+ Proofview.V82.of_tactic (simplest_case v);
+ Proofview.V82.of_tactic intros;
observe_tac "reflexivity_with_destruct_cases" reflexivity_with_destruct_cases
]
- | _ -> reflexivity
- with e when Errors.noncritical e -> reflexivity
+ | _ -> Proofview.V82.of_tactic reflexivity
+ with e when Errors.noncritical e -> Proofview.V82.of_tactic reflexivity
in
- let eq_ind = Coqlib.build_coq_eq () in
+ let eq_ind = make_eq () in
let discr_inject =
Tacticals.onAllHypsAndConcl (
fun sc g ->
@@ -586,15 +854,15 @@ let rec reflexivity_with_destruct_cases g =
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
+ then Proofview.V82.of_tactic (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
+ then tclTHENSEQ [Proofview.V82.of_tactic (Equality.injHyp None id);thin [id];intros_with_rewrite] g
else tclIDTAC g
| _ -> tclIDTAC g
)
in
(tclFIRST
- [ observe_tac "reflexivity_with_destruct_cases : reflexivity" reflexivity;
+ [ observe_tac "reflexivity_with_destruct_cases : reflexivity" (Proofview.V82.of_tactic reflexivity);
observe_tac "reflexivity_with_destruct_cases : destruct_case" ((destruct_case ()));
(* We reach this point ONLY if
the same value is matched (at least) two times
@@ -654,23 +922,24 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
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 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
+ 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 *)
+ (* we also compute fresh names for each hyptohesis of each branch
+ 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))
+ (generate_fresh_id (Id.of_string "y") ids (nb_prod br_type))
)
branches
in
@@ -680,28 +949,34 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
*)
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
+ let infos =
+ try find_Function_infos (fst (destConst funcs.(j)))
+ with Not_found -> error "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 "Cannot find equation lemma"
+ with Option.IsNone -> anomaly (Pp.str "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
- (Glob_term.Cbv
- {Glob_term.all_flags
- with Glob_term.rDelta = false;
+ 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
+ have been $\zeta$-normalized *)
+ reduce
+ (Genredexpr.Cbv
+ {Redops.all_flags
+ with Genredexpr.rDelta = false;
})
- onConcl
+ Locusops.onConcl
;
- h_generalize (List.map mkVar ids);
+ Simple.generalize (List.map mkVar ids);
thin ids
]
- else unfold_in_concl [(Termops.all_occurrences, Names.EvalConstRef (destConst f))]
+ else
+ unfold_in_concl [(Locus.AllOccurrences, Names.EvalConstRef (fst (destConst f)))]
in
(* The proof of each branche itself *)
let ind_number = ref 0 in
@@ -725,21 +1000,21 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
(* 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;
+ observe_tac "intros_with_rewrite (all)" 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_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]);
+ [ tclMAP (fun id -> Proofview.V82.of_tactic (Simple.intro id)) (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;
+ (Simple.generalize [mkApp(applist(graph_principle,params),Array.map (fun c -> applist(c,params)) lemmas)]);
+ Proofview.V82.of_tactic (Simple.intro graph_principle_id);
observe_tac "" (tclTHEN_i
- (observe_tac "elim" ((elim false (mkVar hres,Glob_term.NoBindings) (Some (mkVar graph_principle_id,Glob_term.NoBindings)))))
+ (observe_tac "elim" (Proofview.V82.of_tactic (elim false None (mkVar hres,NoBindings) (Some (mkVar graph_principle_id,NoBindings)))))
(fun i g -> observe_tac "prove_branche" (prove_branche i) g ))
]
g
@@ -747,7 +1022,7 @@ let prove_fun_complete funcs graphs schemes lemmas_types_infos i : tactic =
-let do_save () = Lemmas.save_named false
+let do_save () = Lemmas.save_proof (Vernacexpr.Proved(false,None))
(* [derive_correctness make_scheme functional_induction funs graphs] create correctness and completeness
@@ -758,15 +1033,14 @@ let do_save () = Lemmas.save_named false
*)
let derive_correctness make_scheme functional_induction (funs: constant list) (graphs:inductive list) =
- let previous_state = States.freeze () in
let funs = Array.of_list funs and graphs = Array.of_list graphs in
let funs_constr = Array.map mkConst funs in
- try
+ States.with_state_protection_on_exception (fun () ->
let graphs_constr = Array.map mkInd graphs in
let lemmas_types_infos =
- Util.array_map2_i
+ Util.Array.map2_i
(fun i f_constr graph ->
- let const_of_f = destConst f_constr in
+ let const_of_f,u = 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
@@ -783,15 +1057,15 @@ let derive_correctness make_scheme functional_induction (funs: constant list) (g
if the block contains only one function we can safely reuse [f_rect]
*)
try
- if Array.length funs_constr <> 1 then raise Not_found;
+ if not (Int.equal (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 )
+ (fst (fst(Future.force entry.Entries.const_entry_body)), Option.get entry.Entries.const_entry_type )
)
- (make_scheme (array_map_to_list (fun const -> const,Glob_term.GType None) funs))
+ (make_scheme (Array.map_to_list (fun const -> const,GType []) funs))
)
in
let proving_tac =
@@ -799,28 +1073,29 @@ let derive_correctness make_scheme functional_induction (funs: constant list) (g
in
Array.iteri
(fun i f_as_constant ->
- let f_id = id_of_label (con_label f_as_constant) in
+ let f_id = Label.to_id (con_label f_as_constant) in
(*i The next call to mk_correct_id is valid since we are constructing the lemma
Ensures by: obvious
i*)
let lem_id = mk_correct_id f_id in
Lemmas.start_proof lem_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));
+ (Decl_kinds.Global,false(*FIXME*),(Decl_kinds.Proof Decl_kinds.Theorem))
+ (*FIXME*) Evd.empty
+ (fst lemmas_types_infos.(i))
+ (Lemmas.mk_hook (fun _ _ -> ()));
+ ignore (Pfedit.by
+ (Proofview.V82.tactic (observe_tac ("prove correctness ("^(Id.to_string f_id)^")")
+ (proving_tac i))));
do_save ();
let finfo = find_Function_infos f_as_constant in
- let lem_cst = destConst (Constrintern.global_reference lem_id) in
+ let lem_cst = fst (destConst (Constrintern.global_reference lem_id)) in
update_Function {finfo with correctness_lemma = Some lem_cst}
)
funs;
let lemmas_types_infos =
- Util.array_map2_i
+ Util.Array.map2_i
(fun i f_constr graph ->
- let const_of_f = destConst f_constr in
+ let const_of_f = fst (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
@@ -832,51 +1107,46 @@ let derive_correctness make_scheme functional_induction (funs: constant list) (g
funs_constr
graphs_constr
in
- let kn,_ as graph_ind = destInd graphs_constr.(0) in
+ let kn,_ as graph_ind = fst (destInd graphs_constr.(0)) in
let mib,mip = Global.lookup_inductive graph_ind in
- let schemes =
- Array.of_list
+ let sigma, scheme =
(Indrec.build_mutual_induction_scheme (Global.env ()) Evd.empty
(Array.to_list
(Array.mapi
- (fun i _ -> (kn,i),true,InType)
+ (fun i _ -> ((kn,i),Univ.Instance.empty)(*FIXME*),true,InType)
mib.Declarations.mind_packets
)
)
)
in
+ let schemes =
+ Array.of_list scheme
+ 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
+ let f_id = Label.to_id (con_label f_as_constant) in
(*i The next call to mk_complete_id is valid since we are constructing the lemma
Ensures by: obvious
i*)
let lem_id = mk_complete_id f_id in
Lemmas.start_proof lem_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));
+ (Decl_kinds.Global,false(*FIXME*),(Decl_kinds.Proof Decl_kinds.Theorem))
+ (*FIXME*) Evd.empty
+ (fst lemmas_types_infos.(i))
+ (Lemmas.mk_hook (fun _ _ -> ()));
+ ignore (Pfedit.by
+ (Proofview.V82.tactic (observe_tac ("prove completeness ("^(Id.to_string f_id)^")")
+ (proving_tac i))));
do_save ();
let finfo = find_Function_infos f_as_constant in
- let lem_cst = destConst (Constrintern.global_reference lem_id) in
+ let lem_cst,u = destConst (Constrintern.global_reference lem_id) in
update_Function {finfo with completeness_lemma = Some lem_cst}
)
- funs;
- with reraise ->
- (* In case of problem, we reset all the lemmas *)
- Pfedit.delete_all_proofs ();
- States.unfreeze previous_state;
- raise reraise
-
-
-
-
+ funs)
+ ()
(***********************************************)
@@ -890,13 +1160,13 @@ 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'
+ let ((kn',num) as ind'),u = destInd 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 "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
@@ -904,12 +1174,12 @@ let revert_graph kn post_tac hid g =
match info.completeness_lemma with
| None -> tclIDTAC g
| Some f_complete ->
- let f_args,res = array_chop (Array.length args - 1) args in
+ 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])];
+ Simple.generalize [applist(mkConst f_complete,(Array.to_list f_args)@[res.(0);mkVar hid])];
thin [hid];
- h_intro hid;
+ Proofview.V82.of_tactic (Simple.intro hid);
post_tac hid
]
g
@@ -937,26 +1207,26 @@ let revert_graph kn post_tac hid g =
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 old_ids = List.fold_right Id.Set.add (pf_ids_of_hyps g) Id.Set.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 ()) ->
+ | App(eq,args) when eq_constr 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 ->
- ((fun hid -> h_symmetry (onHyp hid)),f_args,args.(2))
+ ((fun hid -> Proofview.V82.of_tactic (intros_symmetry (Locusops.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])];
+ Simple.generalize [applist(f_correct,(Array.to_list f_args)@[res;mkVar hid])];
thin [hid];
- h_intro hid;
- Inv.inv FullInversion None (Glob_term.NamedHyp hid);
+ Proofview.V82.of_tactic (Simple.intro hid);
+ Proofview.V82.of_tactic (Inv.inv FullInversion None (NamedHyp hid));
(fun g ->
- let new_ids = List.filter (fun id -> not (Idset.mem id old_ids)) (pf_ids_of_hyps g) in
+ let new_ids = List.filter (fun id -> not (Id.Set.mem id old_ids)) (pf_ids_of_hyps g) in
tclMAP (revert_graph kn pre_tac) (hid::new_ids) g
);
] g
@@ -968,14 +1238,16 @@ let invfun qhyp f =
let f =
match f with
| ConstRef f -> f
- | _ -> raise (Util.UserError("",str "Not a function"))
+ | _ -> raise (Errors.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
+ Proofview.V82.of_tactic (
+ Tactics.try_intros_until (fun hid -> Proofview.V82.tactic (functional_inversion kn hid (mkConst f) f_correct)) qhyp
+ )
with
| Not_found -> error "No graph found"
| Option.IsNone -> error "Cannot use equivalence with graph!"
@@ -985,16 +1257,17 @@ let invfun qhyp f g =
match f with
| Some f -> invfun qhyp f g
| None ->
+ Proofview.V82.of_tactic begin
Tactics.try_intros_until
- (fun hid g ->
+ (fun hid -> Proofview.V82.tactic begin fun 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 ()) ->
+ | App(eq,args) when eq_constr eq (make_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 finfos = find_Function_infos (fst (destConst f1)) in
let f_correct = mkConst(Option.get finfos.correctness_lemma)
and kn = fst finfos.graph_ind
in
@@ -1003,14 +1276,14 @@ let invfun qhyp f g =
try
let f2,_ = decompose_app args.(2) in
if not (isConst f2) then failwith "";
- let finfos = find_Function_infos (destConst f2) in
+ let finfos = find_Function_infos (fst (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")
+ errorlabstrm "" (str "Hypothesis " ++ Ppconstr.pr_id hid ++ str " must contain at least one Function")
| Option.IsNone ->
if do_observe ()
then
@@ -1023,6 +1296,7 @@ let invfun qhyp f g =
else errorlabstrm "" (str "Cannot find inversion information for hypothesis " ++ Ppconstr.pr_id hid)
end
| _ -> errorlabstrm "" (Ppconstr.pr_id hid ++ str " must be an equality ")
- )
+ end)
qhyp
+ end
g
diff --git a/plugins/funind/merge.ml b/plugins/funind/merge.ml
index e1f10be2..ea699580 100644
--- a/plugins/funind/merge.ml
+++ b/plugins/funind/merge.ml
@@ -1,6 +1,6 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
@@ -8,20 +8,23 @@
(* Merging of induction principles. *)
-open Libnames
+open Globnames
open Tactics
open Indfun_common
+open Errors
open Util
-open Topconstr
+open Constrexpr
open Vernacexpr
open Pp
open Names
open Term
+open Vars
+open Context
open Termops
open Declarations
-open Environ
open Glob_term
open Glob_termops
+open Decl_kinds
(** {1 Utilities} *)
@@ -48,33 +51,33 @@ let rec substitterm prof t by_t 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())
+let understand = Pretyping.understand (Global.env()) Evd.empty
(** Operations on names and identifiers *)
let id_of_name = function
- Anonymous -> id_of_string "H"
+ 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)
+let name_of_string str = Name (Id.of_string str)
+let string_of_name nme = Id.to_string (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
- | GVar (_,x) -> Pervasives.compare x f = 0
+ | GVar (_,x) -> Id.equal x f
| _ -> 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
+ let ans = CRef (Libnames.Ident (Loc.ghost,id), None) in
+ let _ = ignore (Constrintern.intern_constr (Global.env()) ans) in
true
with e when Errors.noncritical e -> 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 next_ident_fresh (id:Id.t) =
let res = ref id in
while ident_global_exist !res do res := Nameops.lift_subscript !res done;
!res
@@ -128,19 +131,15 @@ let prNamedRLDecl s lc =
prstr "\n";
end
-let showind (id:identifier) =
+let showind (id:Id.t) =
let cstrid = Constrintern.global_reference 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
+ let mib1,ib1 = Inductive.lookup_mind_specif (Global.env()) (fst 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?");
+ Printf.printf "arity :"; prconstr (Inductiveops.type_of_inductive (Global.env ()) ind1);
Array.iteri
(fun i x -> Printf.printf"type constr %d :" i ; prconstr x)
ib1.mind_user_lc
@@ -152,23 +151,15 @@ 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
+match Array.findi pred arr with
+| None -> Array.length arr (* all elt are positive *)
+| Some i -> i
-(* Like list_chop but except that [i] is the size of the suffix of [l]. *)
+(* 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
+ 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
@@ -234,7 +225,7 @@ let linkmonad f lnkvar =
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)
+module Link = Map.Make (Int)
let pr_links l =
Printf.printf "links:\n";
@@ -254,7 +245,7 @@ type 'a merged_arg =
type merge_infos =
{
- ident:identifier; (** new inductive name *)
+ ident:Id.t; (** new inductive name *)
mib1: mutual_inductive_body;
oib1: one_inductive_body;
mib2: mutual_inductive_body;
@@ -357,17 +348,17 @@ let filter_shift_stable_right (lnk:int merged_arg array) (l:'a list): 'a list =
(** {1 Utilities for merging} *)
-let ind1name = id_of_string "__ind1"
-let ind2name = id_of_string "__ind2"
+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";
+ if mib1.mind_finite == Decl_kinds.CoFinite then error "First argument is coinductive";
+ if mib2.mind_finite == Decl_kinds.CoFinite then error "Second argument is coinductive";
+ if not (Int.equal mib1.mind_ntypes 1) then error "First argument is mutual";
+ if not (Int.equal mib2.mind_ntypes 1) then error "Second argument is mutual";
()
(*
@@ -381,11 +372,11 @@ let build_raw_params prms_decl avoid =
let _ = prNamedRConstr "RAWDUMMY" dummy_glob_constr in
let res,_ = glob_decompose_prod dummy_glob_constr in
let comblist = List.combine prms_decl res in
- comblist, res , (avoid @ (Idset.elements (ids_of_glob_constr dummy_glob_constr)))
+ comblist, res , (avoid @ (Id.Set.elements (ids_of_glob_constr dummy_glob_constr)))
*)
let ids_of_rawlist avoid rawl =
- List.fold_left Idset.union avoid (List.map ids_of_glob_constr rawl)
+ List.fold_left Id.Set.union avoid (List.map ids_of_glob_constr rawl)
@@ -463,7 +454,7 @@ let shift_linked_params mib1 mib2 (lnk1:linked_var array) (lnk2:linked_var array
([],[],[],[]) arity_ctxt in
(* let arity_ctxt2 =
build_raw_params oib2.mind_arity_ctxt
- (Idset.elements (ids_of_glob_constr oib1.mind_arity_ctxt)) in*)
+ (Id.Set.elements (ids_of_glob_constr 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
@@ -514,16 +505,16 @@ let rec merge_app c1 c2 id1 id2 shift filter_shift_stable =
| GApp(_,f1, arr1), GApp(_,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
- GApp (dummy_loc,GVar (dummy_loc,shift.ident) , args)
+ GApp (Loc.ghost,GVar (Loc.ghost,shift.ident) , args)
| GApp(_,f1, arr1), GApp(_,f2,arr2) -> raise NoMerge
| GLetIn(_,nme,bdy,trm) , _ ->
let _ = prstr "\nICI2!\n";Pp.flush_all() in
let newtrm = merge_app trm c2 id1 id2 shift filter_shift_stable in
- GLetIn(dummy_loc,nme,bdy,newtrm)
+ GLetIn(Loc.ghost,nme,bdy,newtrm)
| _, GLetIn(_,nme,bdy,trm) ->
let _ = prstr "\nICI3!\n";Pp.flush_all() in
let newtrm = merge_app c1 trm id1 id2 shift filter_shift_stable in
- GLetIn(dummy_loc,nme,bdy,newtrm)
+ GLetIn(Loc.ghost,nme,bdy,newtrm)
| _ -> let _ = prstr "\nICI4!\n";Pp.flush_all() in
raise NoMerge
@@ -532,16 +523,16 @@ let rec merge_app_unsafe c1 c2 shift filter_shift_stable =
match c1 , c2 with
| GApp(_,f1, arr1), GApp(_,f2,arr2) ->
let args = filter_shift_stable lnk (arr1 @ arr2) in
- GApp (dummy_loc,GVar(dummy_loc,shift.ident) , args)
+ GApp (Loc.ghost,GVar(Loc.ghost,shift.ident) , args)
(* FIXME: what if the function appears in the body of the let? *)
| GLetIn(_,nme,bdy,trm) , _ ->
let _ = prstr "\nICI2 '!\n";Pp.flush_all() in
let newtrm = merge_app_unsafe trm c2 shift filter_shift_stable in
- GLetIn(dummy_loc,nme,bdy,newtrm)
+ GLetIn(Loc.ghost,nme,bdy,newtrm)
| _, GLetIn(_,nme,bdy,trm) ->
let _ = prstr "\nICI3 '!\n";Pp.flush_all() in
let newtrm = merge_app_unsafe c1 trm shift filter_shift_stable in
- GLetIn(dummy_loc,nme,bdy,newtrm)
+ GLetIn(Loc.ghost,nme,bdy,newtrm)
| _ -> let _ = prstr "\nICI4 '!\n";Pp.flush_all() in raise NoMerge
@@ -550,8 +541,8 @@ let rec merge_app_unsafe c1 c2 shift filter_shift_stable =
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 * glob_constr option * glob_constr option) list)
- filter_shift_stable : (Names.name * glob_constr option * glob_constr option) list =
+ (ltyp:(Name.t * glob_constr option * glob_constr option) list)
+ filter_shift_stable : (Name.t * glob_constr option * glob_constr option) list =
let mergeonehyp t reldecl =
match reldecl with
| (nme,x,Some (GApp(_,i,args) as ind))
@@ -567,11 +558,11 @@ let rec merge_rec_hyps shift accrec
| e::lt -> e :: merge_rec_hyps shift accrec lt filter_shift_stable
-let rec build_suppl_reccall (accrec:(name * glob_constr) list) concl2 shift =
+let build_suppl_reccall (accrec:(Name.t * glob_constr) list) concl2 shift =
List.map (fun (nm,tp) -> (nm,merge_app_unsafe tp concl2 shift)) accrec
-let find_app (nme:identifier) ltyp =
+let find_app (nme:Id.t) ltyp =
try
ignore
(List.map
@@ -591,9 +582,9 @@ let prnt_prod_or_letin nm letbdy typ =
let rec merge_types shift accrec1
- (ltyp1:(name * glob_constr option * glob_constr option) list)
- (concl1:glob_constr) (ltyp2:(name * glob_constr option * glob_constr option) list) concl2
- : (name * glob_constr option * glob_constr option) list * glob_constr =
+ (ltyp1:(Name.t * glob_constr option * glob_constr option) list)
+ (concl1:glob_constr) (ltyp2:(Name.t * glob_constr option * glob_constr option) list) concl2
+ : (Name.t * glob_constr option * glob_constr option) list * glob_constr =
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
@@ -603,7 +594,7 @@ let rec merge_types shift accrec1
let res =
match ltyp1 with
| [] ->
- let isrec1 = (accrec1<>[]) in
+ let isrec1 = not (List.is_empty accrec1) in
let isrec2 = find_app ind2name ltyp2 in
let rechyps =
if isrec1 && isrec2
@@ -657,22 +648,22 @@ let rec merge_types shift accrec1
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)
+let build_link_map_aux (allargs1:Id.t array) (allargs2:Id.t array)
(lnk:int merged_arg array) =
- array_fold_lefti
+ Array.fold_left_i
(fun i acc e ->
- if i = Array.length lnk - 1 then acc (* functional arg, not in allargs *)
+ if Int.equal 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
+ | Prm_linked j | Arg_linked j -> Id.Map.add allargs2.(i) allargs1.(j) acc
| _ -> acc)
- Idmap.empty lnk
+ Id.Map.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
+ Array.of_list (List.rev_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
+ Array.of_list (List.rev_map (fun (x,_,_) -> id_of_name x) allargs2) in
build_link_map_aux allargs1 allargs2 lnk
@@ -749,18 +740,18 @@ let fresh_cstror_suffix , cstror_suffix_init =
(** [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)
+let merge_constructor_id id1 id2 shift:Id.t =
+ let id = Id.to_string 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 glob_constr
first, each of them having distinct var names. *)
-let rec merge_constructors (shift:merge_infos) (avoid:Idset.t)
- (typcstr1:(identifier * glob_constr) list)
- (typcstr2:(identifier * glob_constr) list) : (identifier * glob_constr) list =
+let merge_constructors (shift:merge_infos) (avoid:Id.Set.t)
+ (typcstr1:(Id.t * glob_constr) list)
+ (typcstr2:(Id.t * glob_constr) list) : (Id.t * glob_constr) list =
List.flatten
(List.map
(fun (id1,rawtyp1) ->
@@ -776,20 +767,20 @@ let rec merge_constructors (shift:merge_infos) (avoid:Idset.t)
(** [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)
+let merge_inductive_body (shift:merge_infos) avoid (oib1:one_inductive_body)
(oib2:one_inductive_body) =
(* building glob_constr 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
+ Detyping.detype false (Id.Set.elements avoid) (Global.env()) Evd.empty substindtyp in
let lcstr1: glob_constr 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 avoid2 = Id.Set.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 avoid3 = Id.Set.union avoid (ids_of_rawlist avoid lcstr2) in
let params1 =
try fst (glob_decompose_prod_n shift.nrecprms1 (List.hd lcstr1))
@@ -810,14 +801,14 @@ let rec merge_inductive_body (shift:merge_infos) avoid (oib1:one_inductive_body)
[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
+let 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)
+ merge_inductive_body shift Id.Set.empty mib1.mind_packets.(0) mib2.mind_packets.(0)
let glob_constr_to_constr_expr x = (* build a constr_expr from a glob_constr *)
- Flags.with_option Flags.raw_print (Constrextern.extern_glob_type Idset.empty) x
+ Flags.with_option Flags.raw_print (Constrextern.extern_glob_type Id.Set.empty) x
let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
let params = prms2 @ prms1 in
@@ -828,15 +819,15 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
let _ = prNamedRConstr (string_of_name nme) tp in
let _ = prstr " ; " in
let typ = glob_constr_to_constr_expr tp in
- LocalRawAssum ([(dummy_loc,nme)], Topconstr.default_binder_kind, typ) :: acc)
+ LocalRawAssum ([(Loc.ghost,nme)], Constrexpr_ops.default_binder_kind, typ) :: acc)
[] params in
- let concl = Constrextern.extern_constr false (Global.env()) concl in
+ let concl = Constrextern.extern_constr false (Global.env()) Evd.empty concl in
let arity,_ =
List.fold_left
(fun (acc,env) (nm,_,c) ->
- let typ = Constrextern.extern_constr false env c in
+ let typ = Constrextern.extern_constr false env Evd.empty 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)
+ CProdN (Loc.ghost, [[(Loc.ghost,nm)],Constrexpr_ops.default_binder_kind,typ] , acc) , newenv)
(concl,Global.env())
(shift.funresprms2 @ shift.funresprms1
@ shift.args2 @ shift.args1 @ shift.otherprms2 @ shift.otherprms1) in
@@ -849,33 +840,22 @@ let merge_rec_params_and_arity prms1 prms2 shift (concl:constr) =
[rawlist], named ident.
FIXME: params et cstr_expr (arity) *)
let glob_constr_list_to_inductive_expr prms1 prms2 mib1 mib2 shift
- (rawlist:(identifier * glob_constr) list) =
- let lident = dummy_loc, shift.ident in
+ (rawlist:(Id.t * glob_constr) list) =
+ let lident = Loc.ghost, 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),glob_constr_to_constr_expr t))
+ (fun (id,t) -> false, ((Loc.ghost,id),glob_constr_to_constr_expr t))
rawlist in
lident , bindlist , Some cstr_expr , lcstor_expr
-
-let mkProd_reldecl (rdecl:rel_declaration) (t2:glob_constr) =
- match rdecl with
- | (nme,None,t) ->
- let traw = Detyping.detype false [] [] t in
- GProd (dummy_loc,nme,Explicit,traw,t2)
- | (_,Some _,_) -> assert false
-
-
-
-
let mkProd_reldecl (rdecl:rel_declaration) (t2:glob_constr) =
match rdecl with
| (nme,None,t) ->
- let traw = Detyping.detype false [] [] t in
- GProd (dummy_loc,nme,Explicit,traw,t2)
+ let traw = Detyping.detype false [] (Global.env()) Evd.empty t in
+ GProd (Loc.ghost,nme,Explicit,traw,t2)
| (_,Some _,_) -> assert false
@@ -893,7 +873,7 @@ let merge_inductive (ind1: inductive) (ind2: inductive)
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
+ List.iter (fun (nm,tp) -> prNamedRConstr (Id.to_string nm) tp;prstr "\n") rawlist in
let _ = prstr "\nend rawlist\n" in
(* FIX: retransformer en constr ici
let shift_prm =
@@ -904,15 +884,16 @@ let merge_inductive (ind1: inductive) (ind2: inductive)
let indexpr = glob_constr_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
+ let mie,impls = Command.interp_mutual_inductive indl []
+ false (*FIXMEnon-poly *) false (* means not private *) Decl_kinds.Finite (* means: not coinductive *) in
(* Declare the mutual inductive block with its associated schemes *)
- ignore (Command.declare_mutual_inductive_with_eliminations Declare.UserVerbose mie impls)
+ ignore (Command.declare_mutual_inductive_with_eliminations mie impls)
(* Find infos on identifier id. *)
-let find_Function_infos_safe (id:identifier): Indfun_common.function_info =
+let find_Function_infos_safe (id:Id.t): Indfun_common.function_info =
let kn_of_id x =
- let f_ref = Libnames.Ident (dummy_loc,x) in
+ let f_ref = Libnames.Ident (Loc.ghost,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)
@@ -927,8 +908,8 @@ let find_Function_infos_safe (id:identifier): Indfun_common.function_info =
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 merge (id1:Id.t) (id2:Id.t) (args1:Id.t array)
+ (args2:Id.t 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) ?? *)
@@ -938,7 +919,7 @@ let merge (id1:identifier) (id2:identifier) (args1:identifier array)
as above: vars may be linked inside args2?? *)
Array.mapi
(fun i c ->
- match array_find_i (fun i x -> x=c) args1 with
+ match Array.findi (fun i x -> Id.equal x c) args1 with
| Some j -> Linked j
| None -> Unlinked)
args2 in
@@ -955,7 +936,7 @@ let remove_last_arg c =
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 rec remove_n_fst_list n l = if Int.equal 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 =
@@ -977,7 +958,7 @@ let funify_branches relinfo nfuns branch =
| _ -> assert false in
let is_dom c =
match kind_of_term c with
- | Ind((u,_)) | Construct((u,_),_) -> u = mut_induct
+ | Ind(((u,_),_)) | Construct(((u,_),_),_) -> MutInd.equal u mut_induct
| _ -> false in
let _dom_i c =
assert (is_dom c);
diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml
index a33ae1d6..5558556e 100644
--- a/plugins/funind/recdef.ml
+++ b/plugins/funind/recdef.ml
@@ -1,59 +1,111 @@
(************************************************************************)
(* v * The Coq Proof Assistant / The Coq Development Team *)
-(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2014 *)
+(* <O___,, * INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2015 *)
(* \VV/ **************************************************************)
(* // * This file is distributed under the terms of the *)
(* * GNU Lesser General Public License Version 2.1 *)
(************************************************************************)
-(*i camlp4deps: "parsing/grammar.cma" i*)
-
open Term
+open Vars
open Namegen
open Environ
-open Declarations
open Entries
open Pp
open Names
open Libnames
+open Globnames
open Nameops
+open Errors
open Util
-open Closure
-open RedFlags
open Tacticals
-open Typing
open Tacmach
open Tactics
open Nametab
-open Decls
open Declare
open Decl_kinds
open Tacred
open Proof_type
-open Vernacinterp
open Pfedit
-open Topconstr
open Glob_term
open Pretyping
-open Pretyping.Default
-open Safe_typing
open Constrintern
-open Hiddentac
+open Misctypes
+open Genredexpr
open Equality
open Auto
open Eauto
-open Genarg
+open Indfun_common
-let compute_renamed_type gls c =
- rename_bound_vars_as_displayed (*no avoid*) [] (*no rels*) []
- (pf_type_of gls c)
-let qed () = Lemmas.save_named true
-let defined () = Lemmas.save_named false
+(* Ugly things which should not be here *)
+
+let coq_constant m s =
+ Coqlib.coq_constant "RecursiveDefinition" m s
+
+let arith_Nat = ["Arith";"PeanoNat";"Nat"]
+let arith_Lt = ["Arith";"Lt"]
+
+let coq_init_constant s =
+ Coqlib.gen_constant_in_modules "RecursiveDefinition" Coqlib.init_modules s
+
+let find_reference sl s =
+ let dp = Names.DirPath.make (List.rev_map Id.of_string sl) in
+ locate (make_qualid dp (Id.of_string s))
+
+let declare_fun f_id kind ?(ctx=Univ.UContext.empty) value =
+ let ce = definition_entry ~univs:ctx value (*FIXME *) in
+ ConstRef(declare_constant f_id (DefinitionEntry ce, kind));;
+
+let defined () = Lemmas.save_proof (Vernacexpr.Proved (false,None))
+
+let def_of_const t =
+ match (kind_of_term t) with
+ Const sp ->
+ (try (match constant_opt_value_in (Global.env ()) sp with
+ | Some c -> c
+ | _ -> raise Not_found)
+ with Not_found ->
+ anomaly (str "Cannot find definition of constant " ++
+ (Id.print (Label.to_id (con_label (fst 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 constr_of_global x =
+ fst (Universes.unsafe_constr_of_global x)
+
+let constant sl s = constr_of_global (find_reference sl s)
+
+let const_of_ref = function
+ ConstRef kn -> kn
+ | _ -> anomaly (Pp.str "ConstRef expected")
+
+
+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
+
+
+
+
+
+(* Generic values *)
let pf_get_new_ids idl g =
let ids = pf_ids_of_hyps g in
List.fold_right
@@ -61,14 +113,98 @@ let pf_get_new_ids idl g =
idl
[]
-let pf_get_new_id id g =
- List.hd (pf_get_new_ids [id] g)
+let compute_renamed_type gls c =
+ rename_bound_vars_as_displayed (*no avoid*) [] (*no rels*) []
+ (pf_type_of gls c)
+let h'_id = Id.of_string "h'"
+let teq_id = Id.of_string "teq"
+let ano_id = Id.of_string "anonymous"
+let x_id = Id.of_string "x"
+let k_id = Id.of_string "k"
+let v_id = Id.of_string "v"
+let def_id = Id.of_string "def"
+let p_id = Id.of_string "p"
+let rec_res_id = Id.of_string "rec_res";;
+let lt = function () -> (coq_init_constant "lt")
+let le = function () -> (coq_init_constant "le")
+let ex = function () -> (coq_init_constant "ex")
+let nat = function () -> (coq_init_constant "nat")
+let iter_ref () =
+ try find_reference ["Recdef"] "iter"
+ with Not_found -> error "module Recdef not loaded"
+let iter = function () -> (constr_of_global (delayed_force iter_ref))
+let eq = function () -> (coq_init_constant "eq")
+let le_lt_SS = function () -> (constant ["Recdef"] "le_lt_SS")
+let le_lt_n_Sm = function () -> (coq_constant arith_Lt "le_lt_n_Sm")
+let le_trans = function () -> (coq_constant arith_Nat "le_trans")
+let le_lt_trans = function () -> (coq_constant arith_Nat "le_lt_trans")
+let lt_S_n = function () -> (coq_constant arith_Lt "lt_S_n")
+let le_n = function () -> (coq_init_constant "le_n")
+let coq_sig_ref = function () -> (find_reference ["Coq";"Init";"Specif"] "sig")
+let coq_O = function () -> (coq_init_constant "O")
+let coq_S = function () -> (coq_init_constant "S")
+let lt_n_O = function () -> (coq_constant arith_Nat "nlt_0_r")
+let max_ref = function () -> (find_reference ["Recdef"] "max")
+let max_constr = function () -> (constr_of_global (delayed_force max_ref))
+let coq_conj = function () -> find_reference Coqlib.logic_module_name "conj"
+
+let f_S t = mkApp(delayed_force coq_S, [|t|]);;
-let h_intros l =
- tclMAP h_intro l
+let rec n_x_id ids n =
+ if Int.equal n 0 then []
+ else let x = next_ident_away_in_goal x_id ids in
+ x::n_x_id (x::ids) (n-1);;
-let debug_queue = Stack.create ()
+let simpl_iter clause =
+ reduce
+ (Lazy
+ {rBeta=true;rIota=true;rZeta= true; rDelta=false;
+ rConst = [ EvalConstRef (const_of_ref (delayed_force iter_ref))]})
+ clause
+
+(* Others ugly things ... *)
+let (value_f:constr list -> global_reference -> constr) =
+ fun al fterm ->
+ let d0 = Loc.ghost 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 context = List.map
+ (fun (x, c) -> Name x, None, c) (List.combine rev_x_id_l (List.rev al))
+ in
+ let env = Environ.push_rel_context context (Global.env ()) in
+ let glob_body =
+ GCases
+ (d0,RegularStyle,None,
+ [GApp(d0, GRef(d0,fterm,None), List.rev_map (fun x_id -> GVar(d0, x_id)) rev_x_id_l),
+ (Anonymous,None)],
+ [d0, [v_id], [PatCstr(d0,(destIndRef
+ (delayed_force coq_sig_ref),1),
+ [PatVar(d0, Name v_id);
+ PatVar(d0, Anonymous)],
+ Anonymous)],
+ GVar(d0,v_id)])
+ in
+ let body = fst (understand env Evd.empty glob_body)(*FIXME*) in
+ it_mkLambda_or_LetIn body context
+
+let (declare_f : Id.t -> 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);;
+
+
+
+(* Debuging mechanism *)
+let debug_queue = Stack.create ()
let rec print_debug_queue b e =
if not (Stack.is_empty debug_queue)
@@ -76,267 +212,445 @@ let rec print_debug_queue b e =
begin
let lmsg,goal = Stack.pop debug_queue in
if b then
- msgnl (lmsg ++ (str " raised exception " ++ Errors.print e) ++ str " on goal " ++ goal)
+ Pp.msg_debug (lmsg ++ (str " raised exception " ++ Errors.print e) ++ str " on goal " ++ goal)
else
begin
- msgnl (str " from " ++ lmsg ++ str " on goal " ++ goal);
+ Pp.msg_debug (str " from " ++ lmsg ++ str " on goal " ++ goal);
end;
print_debug_queue false e;
end
-
+let observe strm =
+ if do_observe ()
+ then Pp.msg_debug strm
+ else ()
+
let do_observe_tac s tac g =
let goal = Printer.pr_goal g in
- let lmsg = (str "recdef : ") ++ (str s) in
+ let lmsg = (str "recdef : ") ++ s in
+ observe (s++fnl());
Stack.push (lmsg,goal) debug_queue;
try
let v = tac g in
ignore(Stack.pop debug_queue);
v
with reraise ->
+ let reraise = Errors.push reraise in
if not (Stack.is_empty debug_queue)
- then
- print_debug_queue true reraise;
- raise reraise
+ then print_debug_queue true (fst (Cerrors.process_vernac_interp_error reraise));
+ iraise reraise
let observe_tac s tac g =
- if Tacinterp.get_debug () <> Tactic_debug.DebugOff
+ if do_observe ()
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);;
+(* Conclusion tactics *)
-let def_of_const t =
- match (kind_of_term t) with
- Const sp ->
- (try (match body_of_constant (Global.lookup_constant sp) with
- | Some c -> Declarations.force c
- | _ -> assert false)
- with e when Errors.noncritical e ->
- anomaly ("Cannot find definition of constant "^
- (string_of_id (id_of_label (con_label sp))))
- )
- |_ -> assert false
+(* 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 -> clear []
+ | Some l -> tclMAP (fun id -> tclTRY (clear [id])) (List.rev l)
+ in
+ tclTHENLIST
+ [
+ clear_tac;
+ if is_mes
+ then tclTHENLIST
+ [
+ unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference
+ (delayed_force Indfun_common.ltof_ref))];
+ tac
+ ]
+ else tac
+ ]
+ g
-let type_of_const t =
- match (kind_of_term t) with
- Const sp -> Typeops.type_of_constant (Global.env()) sp
- |_ -> assert false
+let tclUSER_if_not_mes concl_tac is_mes names_to_suppress =
+ if is_mes
+ then tclCOMPLETE (fun gl -> Proofview.V82.of_tactic (Simple.apply (delayed_force well_founded_ltof)) gl)
+ else (* tclTHEN (Simple.apply (delayed_force acc_intro_generator_function) ) *) (tclUSER concl_tac is_mes names_to_suppress)
+
+
+
+
-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 check_not_nested f t =
- match kind_of_term t with
- | App(g, _) when eq_constr f g ->
- errorlabstrm "recdef" (str "Nested recursive function are not allowed with Function")
- | Var(_) when eq_constr t f -> errorlabstrm "recdef" (str "Nested recursive function are not allowed with Function")
- | _ -> iter_constr (check_not_nested f) t
-
-
-
-
-let rec (find_call_occs : int -> int -> constr -> constr ->
- (constr list -> constr) * constr list list) =
- fun nb_arg nb_lam f expr ->
- match (kind_of_term expr) with
- App (g, args) when eq_constr g f ->
- if Array.length args <> nb_arg then errorlabstrm "recdef" (str "Partial application of function " ++ Printer.pr_lconstr expr ++ str " in its body is not allowed while using Function");
- Array.iter (check_not_nested f) args;
- (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_arg 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_arg 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(_) when eq_constr expr f -> errorlabstrm "recdef" (str "Partial application of function " ++ Printer.pr_lconstr expr ++ str " in its body is not allowed while using Function")
- | Var(id) -> (fun l -> expr), []
- | Meta(_) -> error "Found a metavariable. Can not treat such a term"
- | Evar(_) -> error "Found an evar. Can not treat such a term"
- | Sort(_) -> (fun l -> expr), []
- | Cast(b,_,_) -> find_call_occs nb_arg nb_lam f b
- | Prod(na,t,b) ->
- error "Found a product. Can not treat such a term"
- | Lambda(na,t,b) ->
+(* Travelling term.
+ Both definitions of [f_terminate] and [f_equation] use the same generic
+ travelling mechanism.
+*)
+
+(* [check_not_nested forbidden e] checks that [e] does not contains any variable
+ of [forbidden]
+*)
+let check_not_nested forbidden e =
+ let rec check_not_nested e =
+ match kind_of_term e with
+ | Rel _ -> ()
+ | Var x ->
+ if Id.List.mem x forbidden
+ then error ("check_not_nested : failure "^Id.to_string x)
+ | Meta _ | Evar _ | Sort _ -> ()
+ | Cast(e,_,t) -> check_not_nested e;check_not_nested t
+ | Prod(_,t,b) -> check_not_nested t;check_not_nested b
+ | Lambda(_,t,b) -> check_not_nested t;check_not_nested b
+ | LetIn(_,v,t,b) -> check_not_nested t;check_not_nested b;check_not_nested v
+ | App(f,l) -> check_not_nested f;Array.iter check_not_nested l
+ | Proj (p,c) -> check_not_nested c
+ | Const _ -> ()
+ | Ind _ -> ()
+ | Construct _ -> ()
+ | Case(_,t,e,a) ->
+ check_not_nested t;check_not_nested e;Array.iter check_not_nested a
+ | Fix _ -> error "check_not_nested : Fix"
+ | CoFix _ -> error "check_not_nested : Fix"
+ in
+ try
+ check_not_nested e
+ with UserError(_,p) ->
+ errorlabstrm "_" (str "on expr : " ++ Printer.pr_lconstr e ++ str " " ++ p)
+
+(* ['a info] contains the local information for travelling *)
+type 'a infos =
+ { nb_arg : int; (* function number of arguments *)
+ concl_tac : tactic; (* final tactic to finish proofs *)
+ rec_arg_id : Id.t; (*name of the declared recursive argument *)
+ is_mes : bool; (* type of recursion *)
+ ih : Id.t; (* induction hypothesis name *)
+ f_id : Id.t; (* function name *)
+ f_constr : constr; (* function term *)
+ f_terminate : constr; (* termination proof term *)
+ func : global_reference; (* functionnal reference *)
+ info : 'a;
+ is_main_branch : bool; (* on the main branch or on a matched expression *)
+ is_final : bool; (* final first order term or not *)
+ values_and_bounds : (Id.t*Id.t) list;
+ eqs : Id.t list;
+ forbidden_ids : Id.t list;
+ acc_inv : constr lazy_t;
+ acc_id : Id.t;
+ args_assoc : ((constr list)*constr) list;
+ }
+
+
+type ('a,'b) journey_info_tac =
+ 'a -> (* the arguments of the constructor *)
+ 'b infos -> (* infos of the caller *)
+ ('b infos -> tactic) -> (* the continuation tactic of the caller *)
+ 'b infos -> (* argument of the tactic *)
+ tactic
+
+(* journey_info : specifies the actions to do on the different term constructors during the travelling of the term
+*)
+type journey_info =
+ { letiN : ((Name.t*constr*types*constr),constr) journey_info_tac;
+ lambdA : ((Name.t*types*constr),constr) journey_info_tac;
+ casE : ((constr infos -> tactic) -> constr infos -> tactic) ->
+ ((case_info * constr * constr * constr array),constr) journey_info_tac;
+ otherS : (unit,constr) journey_info_tac;
+ apP : (constr*(constr list),constr) journey_info_tac;
+ app_reC : (constr*(constr list),constr) journey_info_tac;
+ message : string
+ }
+
+
+
+let rec add_vars forbidden e =
+ match kind_of_term e with
+ | Var x -> x::forbidden
+ | _ -> fold_constr add_vars forbidden e
+
+
+let treat_case forbid_new_ids to_intros finalize_tac nb_lam e infos : tactic =
+ fun g ->
+ let rev_context,b = decompose_lam_n nb_lam e in
+ let ids = List.fold_left (fun acc (na,_) ->
+ let pre_id =
+ match na with
+ | Name x -> x
+ | Anonymous -> ano_id
+ in
+ pre_id::acc
+ ) [] rev_context in
+ let rev_ids = pf_get_new_ids (List.rev ids) g in
+ let new_b = substl (List.map mkVar rev_ids) b in
+ tclTHENLIST
+ [
+ h_intros (List.rev rev_ids);
+ Proofview.V82.of_tactic (intro_using teq_id);
+ onLastHypId (fun heq ->
+ tclTHENLIST[
+ thin to_intros;
+ h_intros to_intros;
+ (fun g' ->
+ let ty_teq = pf_type_of g' (mkVar heq) in
+ let teq_lhs,teq_rhs =
+ let _,args = try destApp ty_teq with DestKO -> assert false in
+ args.(1),args.(2)
+ in
+ let new_b' = Termops.replace_term teq_lhs teq_rhs new_b in
+ let new_infos = {
+ infos with
+ info = new_b';
+ eqs = heq::infos.eqs;
+ forbidden_ids =
+ if forbid_new_ids
+ then add_vars infos.forbidden_ids new_b'
+ else infos.forbidden_ids
+ } in
+ finalize_tac new_infos g'
+ )
+ ]
+ )
+ ] g
+
+let rec travel_aux jinfo continuation_tac (expr_info:constr infos) =
+ match kind_of_term expr_info.info with
+ | CoFix _ | Fix _ -> error "Function cannot treat local fixpoint or cofixpoint"
+ | Proj _ -> error "Function cannot treat projections"
+ | LetIn(na,b,t,e) ->
begin
- match find_call_occs nb_arg (succ nb_lam) f b with
- | _, [] -> (* Lambda are authorized as long as they do not contain
- recursives calls *)
- (fun l -> expr),[]
- | _ -> error "Found a lambda which body contains a recursive call. Such terms are not allowed"
+ let new_continuation_tac =
+ jinfo.letiN (na,b,t,e) expr_info continuation_tac
+ in
+ travel jinfo new_continuation_tac
+ {expr_info with info = b; is_final=false}
end
- | LetIn(na,v,t,b) ->
+ | Rel _ -> anomaly (Pp.str "Free var in goal conclusion !")
+ | Prod _ ->
begin
- match find_call_occs nb_arg nb_lam f v, find_call_occs nb_arg (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 "Found a letin with recursive calls in both variable value and body. Such terms are not allowed."
+ try
+ check_not_nested (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
+ jinfo.otherS () expr_info continuation_tac expr_info
+ with e when Errors.noncritical e ->
+ errorlabstrm "Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id)
end
- | Const(_) -> (fun l -> expr), []
- | Ind(_) -> (fun l -> expr), []
- | Construct (_, _) -> (fun l -> expr), []
- | Case(i,t,a,r) ->
- (match find_call_occs nb_arg nb_lam f a with
- cf, (arg1::args) -> (fun l -> mkCase(i, t, (cf l), r)),(arg1::args)
- | _ -> (fun l -> expr),[])
- | Fix(_) -> error "Found a local fixpoint. Can not treat such a term"
- | CoFix(_) -> error "Found a local cofixpoint : 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)));;
+ | Lambda(n,t,b) ->
+ begin
+ try
+ check_not_nested (expr_info.f_id::expr_info.forbidden_ids) expr_info.info;
+ jinfo.otherS () expr_info continuation_tac expr_info
+ with e when Errors.noncritical e ->
+ errorlabstrm "Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain a recursive call to " ++ pr_id expr_info.f_id)
+ end
+ | Case(ci,t,a,l) ->
+ begin
+ let continuation_tac_a =
+ jinfo.casE
+ (travel jinfo) (ci,t,a,l)
+ expr_info continuation_tac in
+ travel
+ jinfo continuation_tac_a
+ {expr_info with info = a; is_main_branch = false;
+ is_final = false}
+ end
+ | App _ ->
+ let f,args = decompose_app expr_info.info in
+ if eq_constr f (expr_info.f_constr)
+ then jinfo.app_reC (f,args) expr_info continuation_tac expr_info
+ else
+ begin
+ match kind_of_term f with
+ | App _ -> assert false (* f is coming from a decompose_app *)
+ | Const _ | Construct _ | Rel _ | Evar _ | Meta _ | Ind _
+ | Sort _ | Prod _ | Var _ ->
+ let new_infos = {expr_info with info=(f,args)} in
+ let new_continuation_tac =
+ jinfo.apP (f,args) expr_info continuation_tac in
+ travel_args jinfo
+ expr_info.is_main_branch new_continuation_tac new_infos
+ | Case _ -> errorlabstrm "Recdef.travel" (str "the term " ++ Printer.pr_lconstr expr_info.info ++ str " can not contain an applied match (See Limitation in Section 2.3 of refman)")
+ | _ -> anomaly (Pp.str "travel_aux : unexpected "++ Printer.pr_lconstr expr_info.info)
+ end
+ | Cast(t,_,_) -> travel jinfo continuation_tac {expr_info with info=t}
+ | Const _ | Var _ | Meta _ | Evar _ | Sort _ | Construct _ | Ind _ ->
+ let new_continuation_tac =
+ jinfo.otherS () expr_info continuation_tac in
+ new_continuation_tac expr_info
+and travel_args jinfo is_final continuation_tac infos =
+ let (f_args',args) = infos.info in
+ match args with
+ | [] ->
+ continuation_tac {infos with info = f_args'; is_final = is_final}
+ | arg::args' ->
+ let new_continuation_tac new_infos =
+ let new_arg = new_infos.info in
+ travel_args jinfo is_final
+ continuation_tac
+ {new_infos with info = (mkApp(f_args',[|new_arg|]),args')}
+ in
+ travel jinfo new_continuation_tac
+ {infos with info=arg;is_final=false}
+and travel jinfo continuation_tac expr_info =
+ observe_tac
+ (str jinfo.message ++ Printer.pr_lconstr expr_info.info)
+ (travel_aux jinfo continuation_tac expr_info)
-let find_reference sl s =
- (locate (make_qualid(Names.make_dirpath
- (List.map id_of_string (List.rev sl)))
- (id_of_string s)));;
+(* Termination proof *)
-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 rec prove_lt hyple g =
+ begin
+ try
+ let (varx,varz) = match decompose_app (pf_concl g) with
+ | _, x::z::_ when isVar x && isVar z -> x, z
+ | _ -> assert false
+ in
+ let h =
+ List.find (fun id ->
+ match decompose_app (pf_type_of g (mkVar id)) with
+ | _, t::_ -> eq_constr t varx
+ | _ -> false
+ ) hyple
+ in
+ let y =
+ List.hd (List.tl (snd (decompose_app (pf_type_of g (mkVar h))))) in
+ tclTHENLIST[
+ Proofview.V82.of_tactic (apply (mkApp(le_lt_trans (),[|varx;y;varz;mkVar h|])));
+ observe_tac (str "prove_lt") (prove_lt hyple)
+ ]
+ with Not_found ->
+ (
+ (
+ tclTHENLIST[
+ Proofview.V82.of_tactic (apply (delayed_force lt_S_n));
+ (observe_tac (str "assumption: " ++ Printer.pr_goal g) (Proofview.V82.of_tactic assumption))
+ ])
+ )
+ end
+ g
+
+let rec destruct_bounds_aux infos (bound,hyple,rechyps) lbounds g =
+ match lbounds with
+ | [] ->
+ 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[
+ Proofview.V82.of_tactic (split (ImplicitBindings [s_max]));
+ Proofview.V82.of_tactic (intro_then
+ (fun id ->
+ Proofview.V82.tactic begin
+ observe_tac (str "destruct_bounds_aux")
+ (tclTHENS (Proofview.V82.of_tactic (simplest_case (mkVar id)))
+ [
+ tclTHENLIST[Proofview.V82.of_tactic (intro_using h_id);
+ Proofview.V82.of_tactic (simplest_elim(mkApp(delayed_force lt_n_O,[|s_max|])));
+ Proofview.V82.of_tactic default_full_auto];
+ tclTHENLIST[
+ observe_tac (str "clearing k ") (clear [id]);
+ h_intros [k;h';def];
+ observe_tac (str "simple_iter") (simpl_iter Locusops.onConcl);
+ observe_tac (str "unfold functional")
+ (unfold_in_concl[(Locus.OnlyOccurrences [1],
+ evaluable_of_global_reference infos.func)]);
+ observe_tac (str "test" ) (
+ tclTHENLIST[
+ list_rewrite true
+ (List.fold_right
+ (fun e acc -> (mkVar e,true)::acc)
+ infos.eqs
+ (List.map (fun e -> (e,true)) rechyps)
+ );
+ (* list_rewrite true *)
+ (* (List.map (fun e -> (mkVar e,true)) infos.eqs) *)
+ (* ; *)
+
+ (observe_tac (str "finishing")
+ (tclORELSE
+ (Proofview.V82.of_tactic intros_reflexivity)
+ (observe_tac (str "calling prove_lt") (prove_lt hyple))))])
+ ]
+ ]
+ )end))
+ ] g
+ | (_,v_bound)::l ->
+ tclTHENLIST[
+ Proofview.V82.of_tactic (simplest_elim (mkVar v_bound));
+ clear [v_bound];
+ tclDO 2 (Proofview.V82.of_tactic intro);
+ onNthHypId 1
+ (fun p_hyp ->
+ (onNthHypId 2
+ (fun p ->
+ tclTHENLIST[
+ Proofview.V82.of_tactic (simplest_elim
+ (mkApp(delayed_force max_constr, [| bound; mkVar p|])));
+ tclDO 3 (Proofview.V82.of_tactic intro);
+ onNLastHypsId 3 (fun lids ->
+ match lids with
+ [hle2;hle1;pmax] ->
+ destruct_bounds_aux infos
+ ((mkVar pmax),
+ hle1::hle2::hyple,(mkVar p_hyp)::rechyps)
+ l
+ | _ -> assert false) ;
+ ]
+ )
+ )
+ )
+ ] g
+
+let destruct_bounds infos =
+ destruct_bounds_aux infos (delayed_force coq_O,[],[]) infos.values_and_bounds
+
+let terminate_app f_and_args expr_info continuation_tac infos =
+ if expr_info.is_final && expr_info.is_main_branch
+ then
+ tclTHENLIST[
+ continuation_tac infos;
+ observe_tac (str "first split")
+ (Proofview.V82.of_tactic (split (ImplicitBindings [infos.info])));
+ observe_tac (str "destruct_bounds (1)") (destruct_bounds infos)
+ ]
+ else continuation_tac infos
+
+let terminate_others _ expr_info continuation_tac infos =
+ if expr_info.is_final && expr_info.is_main_branch
+ then
+ tclTHENLIST[
+ continuation_tac infos;
+ observe_tac (str "first split")
+ (Proofview.V82.of_tactic (split (ImplicitBindings [infos.info])));
+ observe_tac (str "destruct_bounds") (destruct_bounds infos)
+ ]
+ else continuation_tac infos
+
+let terminate_letin (na,b,t,e) expr_info continuation_tac info =
+ let new_e = subst1 info.info e in
+ let new_forbidden =
+ let forbid =
+ try
+ check_not_nested (expr_info.f_id::expr_info.forbidden_ids) b;
+ true
+ with e when Errors.noncritical e -> false
+ in
+ if forbid
+ then
+ match na with
+ | Anonymous -> info.forbidden_ids
+ | Name id -> id::info.forbidden_ids
+ else info.forbidden_ids
+ in
+ continuation_tac {info with info = new_e; forbidden_ids = new_forbidden}
+
+let pf_type c tac gl =
+ let evars, ty = Typing.e_type_of (pf_env gl) (project gl) c in
+ tclTHEN (Refiner.tclEVARS evars) (tac ty) gl
-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);;
+let pf_typel l tac =
+ let rec aux tys l =
+ match l with
+ | [] -> tac (List.rev tys)
+ | hd :: tl -> pf_type hd (fun ty -> aux (ty::tys) tl)
+ in aux [] l
(* This is like the previous one except that it also rewrite on all
hypotheses except the ones given in the first argument. All the
@@ -344,390 +658,355 @@ let mkCaseEq a : tactic =
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 =
+ Id.t list -> constr -> goal sigma -> tactic * Id.t 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
+ Util.List.map_filter
+ (fun (id, _, t) ->
+ if Id.List.mem id not_on_hyp || not (Termops.occur_term expr t)
+ then None else Some 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|])::
+ let new_hyps = mkApp(Lazy.force refl_equal, [|type_of_expr; expr|])::
to_revert_constr in
+ pf_typel new_hyps (fun _ ->
tclTHENLIST
- [h_generalize new_hyps;
+ [Simple.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 Termops.all_occurrences true (* 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 e when Errors.noncritical e ->
- Pp.msgnl (Printer.pr_goal g1 ++ fnl () ++ pr_id teq ++ str ":" ++ Printer.pr_lconstr ty_teq); assert false
- in
- args.(1),args.(2)
- in
- cont_function (mkVar teq::eqs) (Termops.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"
+ Proofview.V82.of_tactic (change_in_concl None
+ (fun sigma ->
+ pattern_occs [Locus.AllOccurrencesBut [1], expr] (pf_env g2) sigma (pf_concl g2))) g2);
+ Proofview.V82.of_tactic (simplest_case expr)]), to_revert
-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)
+let terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos g =
+ let b =
+ try
+ check_not_nested (expr_info.f_id::expr_info.forbidden_ids) a;
+ false
+ with e when Errors.noncritical e ->
+ true
in
- tclTHENSEQ
- [
- clear_tac;
- if is_mes
- then tclTHEN
- (unfold_in_concl [(Termops.all_occurrences, evaluable_of_global_reference
- (delayed_force ltof_ref))])
- tac
- else tac
- ]
+ let a' = infos.info in
+ let new_info =
+ {infos with
+ info = mkCase(ci,t,a',l);
+ is_main_branch = expr_info.is_main_branch;
+ is_final = expr_info.is_final} in
+ let destruct_tac,rev_to_thin_intro =
+ mkDestructEq [expr_info.rec_arg_id] a' g in
+ let to_thin_intro = List.rev rev_to_thin_intro in
+ observe_tac (str "treating case " ++ int (Array.length l) ++ spc () ++ Printer.pr_lconstr a')
+ (try
+ (tclTHENS
+ destruct_tac
+ (List.map_i (fun i e -> observe_tac (str "do treat case") (treat_case b to_thin_intro (next_step continuation_tac) ci.ci_cstr_ndecls.(i) e new_info)) 0 (Array.to_list l)
+ ))
+ with
+ | UserError("Refiner.thensn_tac3",_)
+ | UserError("Refiner.tclFAIL_s",_) ->
+ (observe_tac (str "is computable " ++ Printer.pr_lconstr new_info.info) (next_step continuation_tac {new_info with info = nf_betaiotazeta new_info.info} )
+ ))
g
+
+let terminate_app_rec (f,args) expr_info continuation_tac _ =
+ List.iter (check_not_nested (expr_info.f_id::expr_info.forbidden_ids))
+ args;
+ begin
+ try
+ let v = List.assoc_f (List.equal Constr.equal) args expr_info.args_assoc in
+ let new_infos = {expr_info with info = v} in
+ tclTHENLIST[
+ continuation_tac new_infos;
+ if expr_info.is_final && expr_info.is_main_branch
+ then
+ tclTHENLIST[
+ observe_tac (str "first split")
+ (Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info])));
+ observe_tac (str "destruct_bounds (3)")
+ (destruct_bounds new_infos)
+ ]
+ else
+ tclIDTAC
+ ]
+ with Not_found ->
+ observe_tac (str "terminate_app_rec not found") (tclTHENS
+ (Proofview.V82.of_tactic (simplest_elim (mkApp(mkVar expr_info.ih,Array.of_list args))))
+ [
+ tclTHENLIST[
+ Proofview.V82.of_tactic (intro_using rec_res_id);
+ Proofview.V82.of_tactic intro;
+ onNthHypId 1
+ (fun v_bound ->
+ (onNthHypId 2
+ (fun v ->
+ let new_infos = { expr_info with
+ info = (mkVar v);
+ values_and_bounds =
+ (v,v_bound)::expr_info.values_and_bounds;
+ args_assoc=(args,mkVar v)::expr_info.args_assoc
+ } in
+ tclTHENLIST[
+ continuation_tac new_infos;
+ if expr_info.is_final && expr_info.is_main_branch
+ then
+ tclTHENLIST[
+ observe_tac (str "first split")
+ (Proofview.V82.of_tactic (split (ImplicitBindings [new_infos.info])));
+ observe_tac (str "destruct_bounds (2)")
+ (destruct_bounds new_infos)
+ ]
+ else
+ tclIDTAC
+ ]
+ )
+ )
+ )
+ ];
+ observe_tac (str "proving decreasing") (
+ tclTHENS (* proof of args < formal args *)
+ (Proofview.V82.of_tactic (apply (Lazy.force expr_info.acc_inv)))
+ [
+ observe_tac (str "assumption") (Proofview.V82.of_tactic assumption);
+ tclTHENLIST
+ [
+ tclTRY(list_rewrite true
+ (List.map
+ (fun e -> mkVar e,true)
+ expr_info.eqs
+ )
+ );
+ tclUSER expr_info.concl_tac true
+ (Some (
+ expr_info.ih::expr_info.acc_id::
+ (fun (x,y) -> y)
+ (List.split expr_info.values_and_bounds)
+ )
+ );
+ ]
+ ])
+ ])
+ end
+let terminate_info =
+ { message = "prove_terminate with term ";
+ letiN = terminate_letin;
+ lambdA = (fun _ _ _ _ -> assert false);
+ casE = terminate_case;
+ otherS = terminate_others;
+ apP = terminate_app;
+ app_reC = terminate_app_rec;
+ }
-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 Termops.all_occurrences
- (* dep proofs also: *) true 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 prove_terminate = travel terminate_info
-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 = Termops.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 = Termops.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)
- )
+(* Equation proof *)
+let equation_case next_step (ci,a,t,l) expr_info continuation_tac infos =
+ terminate_case next_step (ci,a,t,l) expr_info continuation_tac infos
-let rec_leaf_terminate nb_arg f_constr concl_tac is_mes acc_inv hrec (func:global_reference) eqs expr =
- match find_call_occs nb_arg 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 rec prove_le g =
+ let x,z =
+ let _,args = decompose_app (pf_concl g) in
+ (List.hd args,List.hd (List.tl args))
+ in
+ tclFIRST[
+ Proofview.V82.of_tactic assumption;
+ Proofview.V82.of_tactic (apply (delayed_force le_n));
+ begin
+ try
+ let matching_fun =
+ pf_is_matching g
+ (Pattern.PApp(Pattern.PRef (reference_of_constr (le ())),[|Pattern.PVar (destVar x);Pattern.PMeta None|])) in
+ let (h,t) = List.find (fun (_,t) -> matching_fun t) (pf_hyps_types g)
+ in
+ let y =
+ let _,args = decompose_app t in
+ List.hd (List.tl args)
+ in
+ tclTHENLIST[
+ Proofview.V82.of_tactic (apply(mkApp(le_trans (),[|x;y;z;mkVar h|])));
+ observe_tac (str "prove_le (rec)") (prove_le)
+ ]
+ with Not_found -> tclFAIL 0 (mt())
+ end;
+ ]
+ g
-let proveterminate nb_arg 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 nb_arg 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_ndecls.(i))
- 0 (Array.to_list l)) g)
- | _, _::_ ->
- (match find_call_occs nb_arg 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 nb_arg 0 f_constr expr with
- _,[] ->
- (try observe_tac "base_leaf" (base_leaf func eqs expr)
- with reraise ->
- (msgerrnl (str "failure in base case");raise reraise ))
- | _, _::_ ->
- observe_tac "rec_leaf"
- (rec_leaf is_mes acc_inv hrec func eqs expr)) in
- v
- with reraise ->
+let rec make_rewrite_list expr_info max = function
+ | [] -> tclIDTAC
+ | (_,p,hp)::l ->
+ observe_tac (str "make_rewrite_list") (tclTHENS
+ (observe_tac (str "rewrite heq on " ++ pr_id p ) (
+ (fun g ->
+ let t_eq = compute_renamed_type g (mkVar hp) in
+ let k,def =
+ 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
+ Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences
+ true (* dep proofs also: *) true
+ (mkVar hp,
+ ExplicitBindings[Loc.ghost,NamedHyp def,
+ expr_info.f_constr;Loc.ghost,NamedHyp k,
+ (f_S max)]) false) g) )
+ )
+ [make_rewrite_list expr_info max l;
+ tclTHENLIST[ (* x < S max proof *)
+ Proofview.V82.of_tactic (apply (delayed_force le_lt_n_Sm));
+ observe_tac (str "prove_le(2)") prove_le
+ ]
+ ] )
+
+let make_rewrite expr_info l hp max =
+ tclTHENFIRST
+ (observe_tac (str "make_rewrite") (make_rewrite_list expr_info max l))
+ (observe_tac (str "make_rewrite") (tclTHENS
+ (fun g ->
+ let t_eq = compute_renamed_type g (mkVar hp) in
+ let k,def =
+ 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
+ observe_tac (str "general_rewrite_bindings")
+ (Proofview.V82.of_tactic (general_rewrite_bindings false Locus.AllOccurrences
+ true (* dep proofs also: *) true
+ (mkVar hp,
+ ExplicitBindings[Loc.ghost,NamedHyp def,
+ expr_info.f_constr;Loc.ghost,NamedHyp k,
+ (f_S (f_S max))]) false)) g)
+ [observe_tac(str "make_rewrite finalize") (
+ (* tclORELSE( h_reflexivity) *)
+ (tclTHENLIST[
+ simpl_iter Locusops.onConcl;
+ observe_tac (str "unfold functional")
+ (unfold_in_concl[(Locus.OnlyOccurrences [1],
+ evaluable_of_global_reference expr_info.func)]);
+
+ (list_rewrite true
+ (List.map (fun e -> mkVar e,true) expr_info.eqs));
+ (observe_tac (str "h_reflexivity") (Proofview.V82.of_tactic intros_reflexivity))]))
+ ;
+ tclTHENLIST[ (* x < S (S max) proof *)
+ Proofview.V82.of_tactic (apply (delayed_force le_lt_SS));
+ observe_tac (str "prove_le (3)") prove_le
+ ]
+ ])
+ )
+
+let rec compute_max rew_tac max l =
+ match l with
+ | [] -> rew_tac max
+ | (_,p,_)::l ->
+ tclTHENLIST[
+ Proofview.V82.of_tactic (simplest_elim
+ (mkApp(delayed_force max_constr, [| max; mkVar p|])));
+ tclDO 3 (Proofview.V82.of_tactic intro);
+ onNLastHypsId 3 (fun lids ->
+ match lids with
+ | [hle2;hle1;pmax] -> compute_max rew_tac (mkVar pmax) l
+ | _ -> assert false
+ )]
+
+let rec destruct_hex expr_info acc l =
+ match l with
+ | [] ->
begin
- msgerrnl(str "failure in proveterminate");
- raise reraise
+ match List.rev acc with
+ | [] -> tclIDTAC
+ | (_,p,hp)::tl ->
+ observe_tac (str "compute max ") (compute_max (make_rewrite expr_info tl hp) (mkVar p) tl)
end
- in
- proveterminate
+ | (v,hex)::l ->
+ tclTHENLIST[
+ Proofview.V82.of_tactic (simplest_case (mkVar hex));
+ clear [hex];
+ tclDO 2 (Proofview.V82.of_tactic intro);
+ onNthHypId 1 (fun hp ->
+ onNthHypId 2 (fun p ->
+ observe_tac
+ (str "destruct_hex after " ++ pr_id hp ++ spc () ++ pr_id p)
+ (destruct_hex expr_info ((v,p,hp)::acc) l)
+ )
+ )
+ ]
+
+let rec intros_values_eq expr_info acc =
+ tclORELSE(
+ tclTHENLIST[
+ tclDO 2 (Proofview.V82.of_tactic intro);
+ onNthHypId 1 (fun hex ->
+ (onNthHypId 2 (fun v -> intros_values_eq expr_info ((v,hex)::acc)))
+ )
+ ])
+ (tclCOMPLETE (
+ destruct_hex expr_info [] acc
+ ))
+
+let equation_others _ expr_info continuation_tac infos =
+ if expr_info.is_final && expr_info.is_main_branch
+ then
+ tclTHEN
+ (continuation_tac infos)
+ (intros_values_eq expr_info [])
+ else continuation_tac infos
+
+let equation_app f_and_args expr_info continuation_tac infos =
+ if expr_info.is_final && expr_info.is_main_branch
+ then intros_values_eq expr_info []
+ else continuation_tac infos
+
+let equation_app_rec (f,args) expr_info continuation_tac info =
+ begin
+ try
+ let v = List.assoc_f (List.equal Constr.equal) args expr_info.args_assoc in
+ let new_infos = {expr_info with info = v} in
+ observe_tac (str "app_rec found") (continuation_tac new_infos)
+ with Not_found ->
+ if expr_info.is_final && expr_info.is_main_branch
+ then
+ tclTHENLIST
+ [ Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args)));
+ continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc};
+ observe_tac (str "app_rec intros_values_eq") (intros_values_eq expr_info [])
+ ]
+ else
+ tclTHENLIST[
+ Proofview.V82.of_tactic (simplest_case (mkApp (expr_info.f_terminate,Array.of_list args)));
+ observe_tac (str "app_rec not_found") (continuation_tac {expr_info with args_assoc = (args,delayed_force coq_O)::expr_info.args_assoc})
+ ]
+ end
-let hyp_terminates nb_args func =
- let a_arrow_b = arg_type (constr_of_global func) in
+let equation_info =
+ {message = "prove_equation with term ";
+ letiN = (fun _ -> assert false);
+ lambdA = (fun _ _ _ _ -> assert false);
+ casE = equation_case;
+ otherS = equation_others;
+ apP = equation_app;
+ app_reC = equation_app_rec
+}
+
+let prove_eq = travel equation_info
+
+(* wrappers *)
+(* [compute_terminate_type] computes the type of the Definition f_terminate from the type of f_F
+*)
+let compute_terminate_type nb_args func =
+ let _,a_arrow_b,_ = destLambda(def_of_const (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)
+ List.rev (List.map_i (fun i _ -> mkRel (6+i)) 0 rev_args)
)
)
in
@@ -744,18 +1023,12 @@ let hyp_terminates nb_args func =
delayed_force nat,
(mkProd (Name k_id, delayed_force nat,
mkArrow cond result))))|])in
- let value = mkApp(delayed_force coq_sig,
+ let value = mkApp(constr_of_global (delayed_force coq_sig_ref),
[|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
@@ -763,14 +1036,14 @@ let termination_proof_header is_mes input_type ids args_id relation
let nargs = List.length args_id in
let pre_rec_args =
List.rev_map
- mkVar (fst (list_chop (rec_arg_num - 1) args_id))
+ 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_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)))
+ (Id.of_string ("Acc_"^(Id.to_string rec_arg_id)))
(wf_thm::ids)
in
let hrec = next_ident_away_in_goal hrec_id
@@ -787,46 +1060,46 @@ let termination_proof_header is_mes input_type ids args_id relation
(h_intros args_id)
(tclTHENS
(observe_tac
- "first assert"
- (assert_tac
+ (str "first assert")
+ (Proofview.V82.of_tactic (assert_before
(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
+ (str "second assert")
+ (Proofview.V82.of_tactic (assert_before
(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));
+ observe_tac (str "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|]))
+ (str "apply wf_thm")
+ (Proofview.V82.of_tactic (Simple.apply (mkApp(mkVar wf_thm,[|mkVar rec_arg_id|])))
)
]
;
(* rest of the proof *)
- tclTHENSEQ
- [observe_tac "generalize"
+ tclTHENLIST
+ [observe_tac (str "generalize")
(onNLastHypsId (nargs+1)
(tclMAP (fun id ->
- tclTHEN (h_generalize [mkVar id]) (h_clear false [id]))
+ tclTHEN (Tactics.Simple.generalize [mkVar id]) (clear [id]))
))
;
- observe_tac "h_fix" (h_fix (Some hrec) (nargs+1));
+ observe_tac (str "fix") (fix (Some hrec) (nargs+1));
h_intros args_id;
- h_intro wf_rec_arg;
- observe_tac "tac" (tac wf_rec_arg hrec acc_inv)
+ Proofview.V82.of_tactic (Simple.intro wf_rec_arg);
+ observe_tac (str "tac") (tac wf_rec_arg hrec wf_rec_arg acc_inv)
]
]
) g
@@ -838,10 +1111,8 @@ let rec instantiate_lambda t l =
match l with
| [] -> t
| a::l ->
- let (bound_name, _, body) = destLambda t in
+ let (_, _, 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
@@ -852,7 +1123,7 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
let f_id =
match f_name with
| Name f_id -> next_ident_away_in_goal f_id ids
- | Anonymous -> anomaly "Anonymous function"
+ | Anonymous -> anomaly (Pp.str "Anonymous function")
in
let n_names_types,_ = decompose_lam_n nb_args body1 in
let n_ids,ids =
@@ -862,7 +1133,7 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
| Name id ->
let n_id = next_ident_away_in_goal id ids in
n_id::n_ids,n_id::ids
- | _ -> anomaly "anonymous argument"
+ | _ -> anomaly (Pp.str "anonymous argument")
)
([],(f_id::ids))
n_names_types
@@ -877,20 +1148,28 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
relation
rec_arg_num
rec_arg_id
- (fun rec_arg_id hrec acc_inv g ->
- (proveterminate
- nb_args
- [rec_arg_id]
- is_mes
- acc_inv
- hrec
- (mkVar f_id)
- func
- base_leaf_terminate
- (rec_leaf_terminate nb_args (mkVar f_id) concl_tac)
- []
- expr
- )
+ (fun rec_arg_id hrec acc_id acc_inv g ->
+ (prove_terminate (fun infos -> tclIDTAC)
+ { is_main_branch = true; (* we are on the main branche (i.e. still on a match ... with .... end *)
+ is_final = true; (* and on leaf (more or less) *)
+ f_terminate = delayed_force coq_O;
+ nb_arg = nb_args;
+ concl_tac = concl_tac;
+ rec_arg_id = rec_arg_id;
+ is_mes = is_mes;
+ ih = hrec;
+ f_id = f_id;
+ f_constr = mkVar f_id;
+ func = func;
+ info = expr;
+ acc_inv = acc_inv;
+ acc_id = acc_id;
+ values_and_bounds = [];
+ eqs = [];
+ forbidden_ids = [];
+ args_assoc = []
+ }
+ )
g
)
(tclUSER_if_not_mes concl_tac)
@@ -900,7 +1179,7 @@ let whole_start (concl_tac:tactic) nb_args is_mes func input_type relation rec_a
let get_current_subgoals_types () =
let p = Proof_global.give_me_the_proof () in
let { Evd.it=sgs ; sigma=sigma } = Proof.V82.subgoals p in
- List.map (Goal.V82.abstract_type sigma) sgs
+ sigma, List.map (Goal.V82.abstract_type sigma) sgs
let build_and_l l =
let and_constr = Coqlib.build_coq_and () in
@@ -913,7 +1192,8 @@ let build_and_l l =
| App(_,_) ->
let (f,_) = decompose_app t in
eq_constr f (well_founded ())
- | _ -> false
+ | _ ->
+ false
in
let compare t1 t2 =
let b1,b2= is_well_founded t1,is_well_founded t2 in
@@ -928,7 +1208,7 @@ let build_and_l l =
let c,tac,nb = f pl in
mk_and p1 c,
tclTHENS
- (apply (constr_of_global conj_constr))
+ (Proofview.V82.of_tactic (apply (constr_of_global conj_constr)))
[tclIDTAC;
tac
],nb+1
@@ -936,12 +1216,12 @@ let build_and_l 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
+ let rec_res_name = Id.to_string rec_res_id in
+ let id_name = Id.to_string id in
try
- String.sub id_name 0 (String.length rec_res_name) = rec_res_name
- with e when Errors.noncritical e -> false
-
+ String.equal (String.sub id_name 0 (String.length rec_res_name)) rec_res_name
+ with Invalid_argument _ -> false
+
let clear_goals =
let rec clear_goal t =
match kind_of_term t with
@@ -957,12 +1237,12 @@ let clear_goals =
let build_new_goal_type () =
- let sub_gls_types = get_current_subgoals_types () in
+ let sigma, 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); *)
+ (* Pp.msgnl (str "sub_gls_types2 := " ++ Pp.prlist_with_sep (fun () -> Pp.fnl () ++ Pp.fnl ()) Printer.pr_lconstr sub_gls_types); *)
let res = build_and_l sub_gls_types in
- res
+ sigma, res
let is_opaque_constant c =
let cb = Global.lookup_constant c in
@@ -971,48 +1251,47 @@ let is_opaque_constant c =
| Declarations.Undef _ -> true
| Declarations.Def _ -> false
-let open_new_goal (build_proof:tactic -> tactic -> unit) using_lemmas ref_ goal_name (gls_type,decompose_and_tac,nb_goal) =
+let open_new_goal build_proof sigma 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")
+ try add_suffix current_proof_name "_subproof"
with e when Errors.noncritical e ->
- anomaly "open_new_goal with an unamed theorem"
+ anomaly (Pp.str "open_new_goal with an unamed theorem")
in
- let sign = initialize_named_context_for_proof () in
let na = next_global_ident_away name [] in
if Termops.occur_existential gls_type then
- Util.error "\"abstract\" cannot handle existentials";
+ Errors.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
+ let na_ref = Libnames.Ident (Loc.ghost,na) in
+ let na_global = Smartlocate.global_with_alias na_ref in
match na_global with
ConstRef c -> is_opaque_constant c
- | _ -> anomaly "equation_lemma: not a constant"
+ | _ -> anomaly ~label:"equation_lemma" (Pp.str "not a constant")
in
- let lemma = mkConst (Lib.make_con na) in
+ let lemma = mkConst (Names.Constant.make1 (Lib.make_kn na)) in
ref_ := Some lemma ;
let lid = ref [] in
let h_num = ref (-1) in
- Flags.silently Vernacentries.interp (Vernacexpr.VernacAbort None);
- build_proof
+ Proof_global.discard_all ();
+ build_proof Evd.empty
( fun gls ->
let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in
- tclTHENSEQ
+ tclTHENLIST
[
- h_generalize [lemma];
- h_intro hid;
+ Simple.generalize [lemma];
+ Proofview.V82.of_tactic (Simple.intro hid);
(fun g ->
let ids = pf_ids_of_hyps g in
tclTHEN
- (Elim.h_decompose_and (mkVar hid))
+ (Proofview.V82.of_tactic (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];
+ lid := List.rev (List.subtract Id.equal ids' ids);
+ if List.is_empty !lid then lid := [hid];
tclIDTAC g
)
g
@@ -1021,40 +1300,39 @@ let open_new_goal (build_proof:tactic -> tactic -> unit) using_lemmas ref_ goal_
(fun g ->
match kind_of_term (pf_concl g) with
| App(f,_) when eq_constr f (well_founded ()) ->
- Auto.h_auto None [] (Some []) g
+ Proofview.V82.of_tactic (Auto.h_auto None [] (Some [])) g
| _ ->
incr h_num;
- (observe_tac "finishing using"
+ (observe_tac (str "finishing using")
(
tclCOMPLETE(
tclFIRST[
tclTHEN
- (eapply_with_bindings (mkVar (List.nth !lid !h_num), NoBindings))
+ (Proofview.V82.of_tactic (eapply_with_bindings (mkVar (List.nth !lid !h_num), NoBindings)))
e_assumption;
Eauto.eauto_with_bases
(true,5)
- [Evd.empty,delayed_force refl_equal]
- [Auto.Hint_db.empty empty_transparent_state false]
+ [Evd.empty,Lazy.force refl_equal]
+ [Hints.Hint_db.empty empty_transparent_state false]
]
)
)
)
g)
;
- Lemmas.save_named opacity;
+ Lemmas.save_proof (Vernacexpr.Proved(opacity,None));
in
- start_proof
+ Lemmas.start_proof
na
- (Decl_kinds.Global, Decl_kinds.Proof Decl_kinds.Lemma)
- sign
- gls_type
- hook ;
+ (Decl_kinds.Global, false (* FIXME *), Decl_kinds.Proof Decl_kinds.Lemma)
+ sigma gls_type
+ (Lemmas.mk_hook hook);
if Indfun_common.is_strict_tcc ()
then
- by (tclIDTAC)
+ ignore (by (Proofview.V82.tactic (tclIDTAC)))
else
begin
- by (
+ ignore (by (Proofview.V82.tactic begin
fun g ->
tclTHEN
(decompose_and_tac)
@@ -1062,23 +1340,21 @@ let open_new_goal (build_proof:tactic -> tactic -> unit) using_lemmas ref_ goal_
(tclFIRST
(List.map
(fun c ->
- tclTHENSEQ
+ Proofview.V82.of_tactic (Tacticals.New.tclTHENLIST
[intros;
- h_simplest_apply (interp_constr Evd.empty (Global.env()) c);
- tclCOMPLETE Auto.default_auto
- ]
+ Simple.apply (fst (interp_constr (Global.env()) Evd.empty c)) (*FIXME*);
+ Tacticals.New.tclCOMPLETE Auto.default_auto
+ ])
)
using_lemmas)
) tclIDTAC)
- g)
+ g end))
end;
try
- by tclIDTAC; (* raises UserError _ if the proof is complete *)
- if Flags.is_verbose () then (pp (Printer.pr_open_subgoals()))
+ ignore (by (Proofview.V82.tactic tclIDTAC)); (* raises UserError _ if the proof is complete *)
with UserError _ ->
defined ()
-;;
let com_terminate
@@ -1090,25 +1366,28 @@ let com_terminate
relation
rec_arg_num
thm_name using_lemmas
- nb_args
+ nb_args ctx
hook =
- let start_proof (tac_start:tactic) (tac_end:tactic) =
+ let start_proof ctx (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;
+ Lemmas.start_proof thm_name
+ (Global, false (* FIXME *), Proof Lemma) ~sign:(Environ.named_context_val env)
+ ctx (compute_terminate_type 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 ))
+ ignore (by (Proofview.V82.tactic (observe_tac (str "starting_tac") tac_start)));
+ ignore (by (Proofview.V82.tactic (observe_tac (str "whole_start") (whole_start tac_end nb_args is_mes fonctional_ref
+ input_type relation rec_arg_num ))))
in
- start_proof tclIDTAC tclIDTAC;
+ start_proof ctx tclIDTAC tclIDTAC;
try
- let new_goal_type = build_new_goal_type () in
- open_new_goal start_proof using_lemmas tcc_lemma_ref
+ let sigma, new_goal_type = build_new_goal_type () in
+ let sigma =
+ Evd.from_env ~ctx:(Evd.evar_universe_context sigma) Environ.empty_env
+ in
+ open_new_goal start_proof sigma
+ 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 ()
@@ -1116,301 +1395,87 @@ let com_terminate
-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 context = List.map
- (fun (x, c) -> Name x, None, c) (List.combine rev_x_id_l (List.rev al))
- in
- let env = Environ.push_rel_context context (Global.env ()) in
- let glob_body =
- GCases
- (d0,RegularStyle,None,
- [GApp(d0, GRef(d0,fterm), List.rev_map (fun x_id -> GVar(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)],
- GVar(d0,v_id)])
- in
- let body = understand Evd.empty env glob_body in
- it_mkLambda_or_LetIn body context
-
-let (declare_fun : identifier -> logical_kind -> constr -> global_reference) =
- fun f_id kind value ->
- let ce = {const_entry_body = value;
- const_entry_secctx = None;
- const_entry_type = None;
- const_entry_opaque = false } 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 =
+ (cont_tactic:Id.t 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 nargs = nb_prod (fst (type_of_const terminate_constr)) (*FIXME*) in
let x = n_x_id ids nargs in
tclTHENLIST [
h_intros x;
- unfold_in_concl [(Termops.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|]);;
+ unfold_in_concl [(Locus.AllOccurrences, evaluable_of_global_reference f)];
+ observe_tac (str "simplest_case")
+ (Proofview.V82.of_tactic (simplest_case (mkApp (terminate_constr,
+ Array.of_list (List.map mkVar x)))));
+ observe_tac (str "prove_eq") (cont_tactic x)] g;;
-
-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, Termops.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 Termops.all_occurrences
- true (* 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 Termops.all_occurrences true (* 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 nb_arg (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 nb_arg 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 nb_arg termine f functional)
- eqs ci.ci_cstr_ndecls.(i))
- 0 (Array.to_list l)) g)
- | _,_::_ ->
- (match find_call_occs nb_arg 0 f expr with
- _,[] -> observe_tac "base_leaf_eq(1)" (base_leaf_eq functional eqs f)
- | fn,args ->
- fun g ->
- let ids = Termops.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 nb_arg 0 f expr with
- _,[] -> observe_tac "base_leaf_eq(2)" ( base_leaf_eq functional eqs f)
- | fn,args ->
- fun g ->
- let ids = Termops.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 : int -> identifier ->
+let (com_eqn : int -> Id.t ->
global_reference -> global_reference -> global_reference
-> constr -> unit) =
fun nb_arg eq_name functional_ref f_ref terminate_ref equation_lemma_type ->
let opacity =
match terminate_ref with
| ConstRef c -> is_opaque_constant c
- | _ -> anomaly "terminate_lemma: not a constant"
+ | _ -> anomaly ~label:"terminate_lemma" (Pp.str "not a constant")
in
let (evmap, env) = Lemmas.get_current_context() in
- let f_constr = (constr_of_global f_ref) in
+ let evmap =
+ Evd.from_env ~ctx:(Evd.evar_universe_context evmap) Environ.empty_env
+ 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
+ (Lemmas.start_proof eq_name (Global, false, Proof Lemma)
+ ~sign:(Environ.named_context_val env)
+ evmap
+ equation_lemma_type
+ (Lemmas.mk_hook (fun _ _ -> ()));
+ ignore (by
+ (Proofview.V82.tactic (start_equation f_ref terminate_ref
(fun x ->
- prove_eq nb_arg
- (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)
- )
+ prove_eq (fun _ -> tclIDTAC)
+ {nb_arg=nb_arg;
+ f_terminate = constr_of_global terminate_ref;
+ f_constr = f_constr;
+ concl_tac = tclIDTAC;
+ func=functional_ref;
+ info=(instantiate_lambda
+ (def_of_const (constr_of_global functional_ref))
+ (f_constr::List.map mkVar x)
+ );
+ is_main_branch = true;
+ is_final = true;
+ values_and_bounds = [];
+ eqs = [];
+ forbidden_ids = [];
+ acc_inv = lazy (assert false);
+ acc_id = Id.of_string "____";
+ args_assoc = [];
+ f_id = Id.of_string "______";
+ rec_arg_id = Id.of_string "______";
+ is_mes = false;
+ ih = Id.of_string "______";
+ }
)
- );
-(* (try Vernacentries.interp (Vernacexpr.VernacShow Vernacexpr.ShowProof) with _ -> ()); *)
+ )));
+ (* (try Vernacentries.interp (Vernacexpr.VernacShow Vernacexpr.ShowProof) with _ -> ()); *)
(* Vernacentries.interp (Vernacexpr.VernacShow Vernacexpr.ShowScript); *)
- Flags.silently (fun () -> Lemmas.save_named opacity) () ;
+ Flags.silently (fun () -> Lemmas.save_proof (Vernacexpr.Proved(opacity,None))) () ;
(* 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 previous_label = Lib.current_command_label () in
- 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
+ let env = Global.env() in
+ let evd = ref (Evd.from_env env) in
+ let function_type = interp_type_evars env evd type_of_f in
+ let env = push_named (function_name,None,function_type) 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
+ let ty = interp_type_evars env evd ~impls:rec_impls eq in
+ let evm, nf = Evarutil.nf_evars_and_universes !evd in
+ let equation_lemma_type = nf_betaiotazeta (nf ty) in
+ let function_type = nf function_type 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
@@ -1430,35 +1495,35 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
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 Decl_kinds.Definition) res in
+ let functional_ref = declare_fun functional_id (IsDefinition Decl_kinds.Definition) ~ctx:(Evd.universe_context evm) 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
+ fst (*FIXME*)(interp_constr
env_with_pre_rec_args
- r
+ Evd.empty
+ 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 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
- let _ = Table.extraction_inline true [Ident (dummy_loc,term_id)] in
-(* message "start second proof"; *)
- let stop = ref false in
- begin
- try com_eqn (List.length res_vars) equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type)
+ let _ = Table.extraction_inline true [Ident (Loc.ghost,term_id)] in
+ (* message "start second proof"; *)
+ let stop =
+ try com_eqn (List.length res_vars) equation_id functional_ref f_ref term_ref (subst_var function_name equation_lemma_type);
+ false
with e when Errors.noncritical e ->
begin
- if Tacinterp.get_debug () <> Tactic_debug.DebugOff
- then pperrnl (str "Cannot create equation Lemma " ++ Errors.print e)
- else anomaly "Cannot create equation Lemma"
+ if do_observe ()
+ then msg_debug (str "Cannot create equation Lemma " ++ Errors.print e)
+ else anomaly (Pp.str "Cannot create equation Lemma")
;
- stop := true;
+ true
end
- end;
- if not !stop
+ in
+ if not stop
then
let eq_ref = Nametab.locate (qualid_of_ident equation_id ) in
let f_ref = destConst (constr_of_global f_ref)
@@ -1471,9 +1536,9 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
spc () ++ str"is defined" )++ fnl () ++
h 1 (Ppconstr.pr_id equation_id ++
spc () ++ str"is defined" )
- )
+ )
in
- try
+ States.with_state_protection_on_exception (fun () ->
com_terminate
tcc_lemma_name
tcc_lemma_constr
@@ -1483,11 +1548,5 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num
term_id
using_lemmas
(List.length res_vars)
- hook
- with reraise ->
- begin
- (try ignore (Backtrack.backto previous_label)
- with e when Errors.noncritical e -> ());
- (* anomaly "Cannot create termination Lemma" *)
- raise reraise
- end
+ evm (Lemmas.mk_hook hook))
+ ()
diff --git a/plugins/funind/recdef.mli b/plugins/funind/recdef.mli
new file mode 100644
index 00000000..f60eedbe
--- /dev/null
+++ b/plugins/funind/recdef.mli
@@ -0,0 +1,20 @@
+
+
+(* val evaluable_of_global_reference : Libnames.global_reference -> Names.evaluable_global_reference *)
+val tclUSER_if_not_mes :
+ Proof_type.tactic ->
+ bool ->
+ Names.Id.t list option ->
+ Proof_type.tactic
+val recursive_definition :
+bool ->
+ Names.Id.t ->
+ Constrintern.internalization_env ->
+ Constrexpr.constr_expr ->
+ Constrexpr.constr_expr ->
+ int -> Constrexpr.constr_expr -> (Term.pconstant ->
+ Term.constr option ref ->
+ Term.pconstant ->
+ Term.pconstant -> int -> Term.types -> int -> Term.constr -> 'a) -> Constrexpr.constr_expr list -> unit
+
+