aboutsummaryrefslogtreecommitdiff
path: root/plugins
diff options
context:
space:
mode:
authorEmilio Jesus Gallego Arias2020-05-13 02:49:27 +0200
committerEmilio Jesus Gallego Arias2020-05-13 02:49:27 +0200
commit4a5944448e31022593df7d6e7e0318cfea92ea98 (patch)
tree643d10e4058af9db9efbf84d7c91c4d8eeb152e1 /plugins
parent0992c2669324a2ab911f5a4c08c27dc97f2bf371 (diff)
parentc5fdfe6ffef15795ecfde8f18f332ddefd35605e (diff)
Merge PR #12307: Cleaning up the legacy proof engine
Reviewed-by: ejgallego
Diffstat (limited to 'plugins')
-rw-r--r--plugins/cc/cctac.ml4
-rw-r--r--plugins/setoid_ring/newring.ml32
2 files changed, 14 insertions, 22 deletions
diff --git a/plugins/cc/cctac.ml b/plugins/cc/cctac.ml
index 0c305d09e8..c485c38009 100644
--- a/plugins/cc/cctac.ml
+++ b/plugins/cc/cctac.ml
@@ -290,7 +290,6 @@ let constr_of_term c = EConstr.of_constr (constr_of_term c)
let rec proof_tac p : unit Proofview.tactic =
Proofview.Goal.enter begin fun gl ->
- try (* type_of can raise exceptions *)
match p.p_rule with
Ax c -> exact_check (EConstr.of_constr c)
| SymAx c ->
@@ -350,7 +349,6 @@ let rec proof_tac p : unit Proofview.tactic =
app_global_with_holes _f_equal [|intype;outtype;proj;ti;tj|] 1 in
Tacticals.New.tclTHEN (Proofview.Unsafe.tclEVARS sigma)
(Tacticals.New.tclTHEN injt (proof_tac prf))))
- with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
end
let refute_tac c t1 t2 p =
@@ -508,11 +506,9 @@ let f_equal =
let concl = Proofview.Goal.concl gl in
let sigma = Tacmach.New.project gl in
let cut_eq c1 c2 =
- try (* type_of can raise an exception *)
Tacticals.New.tclTHENS
(mk_eq _eq c1 c2 Tactics.cut)
[Proofview.tclUNIT ();Tacticals.New.tclTRY ((app_global _refl_equal [||]) apply)]
- with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
in
Proofview.tclORELSE
begin match EConstr.kind sigma concl with
diff --git a/plugins/setoid_ring/newring.ml b/plugins/setoid_ring/newring.ml
index 633cdbd735..e7c75e029e 100644
--- a/plugins/setoid_ring/newring.ml
+++ b/plugins/setoid_ring/newring.ml
@@ -690,15 +690,13 @@ let ring_lookup (f : Value.t) lH rl t =
Proofview.Goal.enter begin fun gl ->
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
- try (* find_ring_strucure can raise an exception *)
- let rl = make_args_list sigma rl t in
- let evdref = ref sigma in
- let e = find_ring_structure env sigma rl in
- let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.ring_carrier) rl) in
- let lH = carg (make_hyp_list env evdref lH) in
- let ring = ltac_ring_structure e in
- Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (Value.apply f (ring@[lH;rl]))
- with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
+ let rl = make_args_list sigma rl t in
+ let evdref = ref sigma in
+ let e = find_ring_structure env sigma rl in
+ let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.ring_carrier) rl) in
+ let lH = carg (make_hyp_list env evdref lH) in
+ let ring = ltac_ring_structure e in
+ Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (Value.apply f (ring@[lH;rl]))
end
(***********************************************************************)
@@ -984,13 +982,11 @@ let field_lookup (f : Value.t) lH rl t =
Proofview.Goal.enter begin fun gl ->
let sigma = Tacmach.New.project gl in
let env = Proofview.Goal.env gl in
- try
- let rl = make_args_list sigma rl t in
- let evdref = ref sigma in
- let e = find_field_structure env sigma rl in
- let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.field_carrier) rl) in
- let lH = carg (make_hyp_list env evdref lH) in
- let field = ltac_field_structure e in
- Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (Value.apply f (field@[lH;rl]))
- with e when Proofview.V82.catchable_exception e -> Proofview.tclZERO e
+ let rl = make_args_list sigma rl t in
+ let evdref = ref sigma in
+ let e = find_field_structure env sigma rl in
+ let rl = Value.of_constr (make_term_list env evdref (EConstr.of_constr e.field_carrier) rl) in
+ let lH = carg (make_hyp_list env evdref lH) in
+ let field = ltac_field_structure e in
+ Proofview.tclTHEN (Proofview.Unsafe.tclEVARS !evdref) (Value.apply f (field@[lH;rl]))
end