aboutsummaryrefslogtreecommitdiffhomepage
path: root/plugins/firstorder
diff options
context:
space:
mode:
authorGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2016-11-26 01:25:11 +0100
committerGravatar Pierre-Marie Pédrot <pierre-marie.pedrot@inria.fr>2017-02-14 17:30:43 +0100
commit01849481fbabc3a3fa6c483e703996b01e37fca5 (patch)
tree80798846a962da7754ddb0b3fae34434de3f7213 /plugins/firstorder
parent8beca748d992cd08e2dd7448c8b28dadbcea4a16 (diff)
Removing compatibility layers from Tacticals
Diffstat (limited to 'plugins/firstorder')
-rw-r--r--plugins/firstorder/instances.ml12
-rw-r--r--plugins/firstorder/rules.ml32
2 files changed, 26 insertions, 18 deletions
diff --git a/plugins/firstorder/instances.ml b/plugins/firstorder/instances.ml
index 2881b5333..ef8172de4 100644
--- a/plugins/firstorder/instances.ml
+++ b/plugins/firstorder/instances.ml
@@ -104,7 +104,7 @@ let mk_open_instance id idc gl m t=
let evmap=Refiner.project gl in
let var_id=
if id==dummy_id then dummy_bvid else
- let typ=pf_unsafe_type_of gl (EConstr.of_constr idc) in
+ let typ=pf_unsafe_type_of gl idc in
(* since we know we will get a product,
reduction is not too expensive *)
let (nam,_,_)=destProd (EConstr.Unsafe.to_constr (whd_all env evmap typ)) in
@@ -127,6 +127,7 @@ let mk_open_instance id idc gl m t=
(* tactics *)
let left_instance_tac (inst,id) continue seq=
+ let open EConstr in
match inst with
Phantom dom->
if lookup (id,None) seq then
@@ -137,8 +138,8 @@ let left_instance_tac (inst,id) continue seq=
[Proofview.V82.of_tactic introf;
pf_constr_of_global id (fun idc ->
(fun gls-> Proofview.V82.of_tactic (generalize
- [EConstr.of_constr (mkApp(idc,
- [|mkVar (Tacmach.pf_nth_hyp_id gls 1)|]))]) gls));
+ [mkApp(idc,
+ [|mkVar (Tacmach.pf_nth_hyp_id gls 1)|])]) gls));
Proofview.V82.of_tactic introf;
tclSOLVE [wrap 1 false continue
(deepen (record (id,None) seq))]];
@@ -152,18 +153,19 @@ let left_instance_tac (inst,id) continue seq=
pf_constr_of_global id (fun idc ->
fun gl->
let evmap,rc,ot = mk_open_instance id idc gl m t in
+ let ot = EConstr.of_constr ot in
let gt=
it_mkLambda_or_LetIn
(mkApp(idc,[|ot|])) rc in
- let gt = EConstr.of_constr gt in
let evmap, _ =
try Typing.type_of (pf_env gl) evmap gt
with e when CErrors.noncritical e ->
error "Untypable instance, maybe higher-order non-prenex quantification" in
tclTHEN (Refiner.tclEVARS evmap) (Proofview.V82.of_tactic (generalize [gt])) gl)
else
+ let t = EConstr.of_constr t in
pf_constr_of_global id (fun idc ->
- Proofview.V82.of_tactic (generalize [EConstr.of_constr (mkApp(idc,[|t|]))]))
+ Proofview.V82.of_tactic (generalize [mkApp(idc,[|t|])]))
in
tclTHENLIST
[special_generalize;
diff --git a/plugins/firstorder/rules.ml b/plugins/firstorder/rules.ml
index 38dae0b20..36bd91ab6 100644
--- a/plugins/firstorder/rules.ml
+++ b/plugins/firstorder/rules.ml
@@ -59,16 +59,17 @@ let clear_global=function
(* connection rules *)
let axiom_tac t seq=
- try pf_constr_of_global (find_left t seq) (fun c -> Proofview.V82.of_tactic (exact_no_check (EConstr.of_constr c)))
+ try pf_constr_of_global (find_left t seq) (fun c -> Proofview.V82.of_tactic (exact_no_check c))
with Not_found->tclFAIL 0 (Pp.str "No axiom link")
let ll_atom_tac a backtrack id continue seq=
+ let open EConstr in
tclIFTHENELSE
(try
tclTHENLIST
[pf_constr_of_global (find_left a seq) (fun left ->
pf_constr_of_global id (fun id ->
- Proofview.V82.of_tactic (generalize [EConstr.of_constr (mkApp(id, [|left|]))])));
+ Proofview.V82.of_tactic (generalize [(mkApp(id, [|left|]))])));
clear_global id;
Proofview.V82.of_tactic intro]
with Not_found->tclFAIL 0 (Pp.str "No link"))
@@ -95,7 +96,7 @@ let left_and_tac ind backtrack id continue seq gls=
let n=(construct_nhyps ind gls).(0) in
tclIFTHENELSE
(tclTHENLIST
- [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim));
+ [Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim);
clear_global id;
tclDO n (Proofview.V82.of_tactic intro)])
(wrap n false continue seq)
@@ -109,12 +110,12 @@ let left_or_tac ind backtrack id continue seq gls=
tclDO n (Proofview.V82.of_tactic intro);
wrap n false continue seq] in
tclIFTHENSVELSE
- (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim)))
+ (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim))
(Array.map f v)
backtrack gls
let left_false_tac id=
- Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim))
+ Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim)
(* left arrow connective rules *)
@@ -133,7 +134,7 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl=
let head=mkApp ((lift p idc),[|capply|]) in
EConstr.of_constr (it_mkLambda_or_LetIn head rc) in
let lp=Array.length rcs in
- let newhyps idc =List.init lp (myterm idc) in
+ let newhyps idc =List.init lp (myterm (EConstr.Unsafe.to_constr idc)) in
tclIFTHENELSE
(tclTHENLIST
[pf_constr_of_global id (fun idc -> Proofview.V82.of_tactic (generalize (newhyps idc)));
@@ -142,17 +143,22 @@ let ll_ind_tac (ind,u as indu) largs backtrack id continue seq gl=
(wrap lp false continue seq) backtrack gl
let ll_arrow_tac a b c backtrack id continue seq=
+ let open EConstr in
+ let open Vars in
+ let a = EConstr.of_constr a in
+ let b = EConstr.of_constr b in
+ let c = EConstr.of_constr c in
let cc=mkProd(Anonymous,a,(lift 1 b)) in
- let d idc =EConstr.of_constr (mkLambda (Anonymous,b,
- mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|]))) in
+ let d idc = mkLambda (Anonymous,b,
+ mkApp (idc, [|mkLambda (Anonymous,(lift 1 a),(mkRel 2))|])) in
tclORELSE
- (tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr c)))
+ (tclTHENS (Proofview.V82.of_tactic (cut c))
[tclTHENLIST
[Proofview.V82.of_tactic introf;
clear_global id;
wrap 1 false continue seq];
- tclTHENS (Proofview.V82.of_tactic (cut (EConstr.of_constr cc)))
- [pf_constr_of_global id (fun c -> Proofview.V82.of_tactic (exact_no_check (EConstr.of_constr c)));
+ tclTHENS (Proofview.V82.of_tactic (cut cc))
+ [pf_constr_of_global id (fun c -> Proofview.V82.of_tactic (exact_no_check c));
tclTHENLIST
[pf_constr_of_global id (fun idc -> Proofview.V82.of_tactic (generalize [d idc]));
clear_global id;
@@ -177,7 +183,7 @@ let forall_tac backtrack continue seq=
let left_exists_tac ind backtrack id continue seq gls=
let n=(construct_nhyps ind gls).(0) in
tclIFTHENELSE
- (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id (EConstr.of_constr %> simplest_elim)))
+ (Proofview.V82.of_tactic (Tacticals.New.pf_constr_of_global id simplest_elim))
(tclTHENLIST [clear_global id;
tclDO n (Proofview.V82.of_tactic intro);
(wrap (n-1) false continue seq)])
@@ -191,9 +197,9 @@ let ll_forall_tac prod backtrack id continue seq=
[Proofview.V82.of_tactic intro;
pf_constr_of_global id (fun idc ->
(fun gls->
+ let open EConstr in
let id0=pf_nth_hyp_id gls 1 in
let term=mkApp(idc,[|mkVar(id0)|]) in
- let term = EConstr.of_constr term in
tclTHEN (Proofview.V82.of_tactic (generalize [term])) (Proofview.V82.of_tactic (clear [id0])) gls));
clear_global id;
Proofview.V82.of_tactic intro;