aboutsummaryrefslogtreecommitdiffhomepage
path: root/plugins/omega/coq_omega.ml
diff options
context:
space:
mode:
authorGravatar Matej Kosik <m4tej.kosik@gmail.com>2016-01-29 10:13:12 +0100
committerGravatar Matej Kosik <m4tej.kosik@gmail.com>2016-02-09 15:58:17 +0100
commit34ef02fac1110673ae74c41c185c228ff7876de2 (patch)
treea688eb9e2c23fc5353391f0c8b4ba1d7ba327844 /plugins/omega/coq_omega.ml
parente9675e068f9e0e92bab05c030fb4722b146123b8 (diff)
CLEANUP: Context.{Rel,Named}.Declaration.t
Originally, rel-context was represented as: Context.rel_context = Names.Name.t * Constr.t option * Constr.t Now it is represented as: Context.Rel.t = LocalAssum of Names.Name.t * Constr.t | LocalDef of Names.Name.t * Constr.t * Constr.t Originally, named-context was represented as: Context.named_context = Names.Id.t * Constr.t option * Constr.t Now it is represented as: Context.Named.t = LocalAssum of Names.Id.t * Constr.t | LocalDef of Names.Id.t * Constr.t * Constr.t Motivation: (1) In "tactics/hipattern.ml4" file we define "test_strict_disjunction" function which looked like this: let test_strict_disjunction n lc = Array.for_all_i (fun i c -> match (prod_assum (snd (decompose_prod_n_assum n c))) with | [_,None,c] -> isRel c && Int.equal (destRel c) (n - i) | _ -> false) 0 lc Suppose that you do not know about rel-context and named-context. (that is the case of people who just started to read the source code) Merlin would tell you that the type of the value you are destructing by "match" is: 'a * 'b option * Constr.t (* worst-case scenario *) or Named.Name.t * Constr.t option * Constr.t (* best-case scenario (?) *) To me, this is akin to wearing an opaque veil. It is hard to figure out the meaning of the values you are looking at. In particular, it is hard to discover the connection between the value we are destructing above and the datatypes and functions defined in the "kernel/context.ml" file. In this case, the connection is there, but it is not visible (between the function above and the "Context" module). ------------------------------------------------------------------------ Now consider, what happens when the reader see the same function presented in the following form: let test_strict_disjunction n lc = Array.for_all_i (fun i c -> match (prod_assum (snd (decompose_prod_n_assum n c))) with | [LocalAssum (_,c)] -> isRel c && Int.equal (destRel c) (n - i) | _ -> false) 0 lc If the reader haven't seen "LocalAssum" before, (s)he can use Merlin to jump to the corresponding definition and learn more. In this case, the connection is there, and it is directly visible (between the function above and the "Context" module). (2) Also, if we already have the concepts such as: - local declaration - local assumption - local definition and we describe these notions meticulously in the Reference Manual, then it is a real pity not to reinforce the connection of the actual code with the abstract description we published.
Diffstat (limited to 'plugins/omega/coq_omega.ml')
-rw-r--r--plugins/omega/coq_omega.ml37
1 files changed, 18 insertions, 19 deletions
diff --git a/plugins/omega/coq_omega.ml b/plugins/omega/coq_omega.ml
index 7e38109d6..b740649e9 100644
--- a/plugins/omega/coq_omega.ml
+++ b/plugins/omega/coq_omega.ml
@@ -28,6 +28,7 @@ open Nametab
open Contradiction
open Misctypes
open Proofview.Notations
+open Context.Named.Declaration
module OmegaSolver = Omega.MakeOmegaSolver (Bigint)
open OmegaSolver
@@ -1695,25 +1696,26 @@ let destructure_hyps =
let pf_nf = Tacmach.New.of_old pf_nf gl in
let rec loop = function
| [] -> (Tacticals.New.tclTHEN nat_inject coq_omega)
- | (i,body,t)::lit ->
+ | decl::lit ->
+ let (i,_,t) = to_tuple decl in
begin try match destructurate_prop t with
| Kapp(False,[]) -> elim_id i
| Kapp((Zle|Zge|Zgt|Zlt|Zne),[t1;t2]) -> loop lit
| Kapp(Or,[t1;t2]) ->
(Tacticals.New.tclTHENS
(elim_id i)
- [ onClearedName i (fun i -> (loop ((i,None,t1)::lit)));
- onClearedName i (fun i -> (loop ((i,None,t2)::lit))) ])
+ [ onClearedName i (fun i -> (loop (LocalAssum (i,t1)::lit)));
+ onClearedName i (fun i -> (loop (LocalAssum (i,t2)::lit))) ])
| Kapp(And,[t1;t2]) ->
Tacticals.New.tclTHEN
(elim_id i)
(onClearedName2 i (fun i1 i2 ->
- loop ((i1,None,t1)::(i2,None,t2)::lit)))
+ loop (LocalAssum (i1,t1) :: LocalAssum (i2,t2) :: lit)))
| Kapp(Iff,[t1;t2]) ->
Tacticals.New.tclTHEN
(elim_id i)
(onClearedName2 i (fun i1 i2 ->
- loop ((i1,None,mkArrow t1 t2)::(i2,None,mkArrow t2 t1)::lit)))
+ loop (LocalAssum (i1,mkArrow t1 t2) :: LocalAssum (i2,mkArrow t2 t1) :: lit)))
| Kimp(t1,t2) ->
(* t1 and t2 might be in Type rather than Prop.
For t1, the decidability check will ensure being Prop. *)
@@ -1724,7 +1726,7 @@ let destructure_hyps =
Proofview.V82.tactic (generalize_tac [mkApp (Lazy.force coq_imp_simp,
[| t1; t2; d1; mkVar i|])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_or (mk_not t1) t2)::lit))))
+ (loop (LocalAssum (i,mk_or (mk_not t1) t2) :: lit))))
]
else
loop lit
@@ -1735,7 +1737,7 @@ let destructure_hyps =
Proofview.V82.tactic (generalize_tac
[mkApp (Lazy.force coq_not_or,[| t1; t2; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_and (mk_not t1) (mk_not t2)):: lit))))
+ (loop (LocalAssum (i,mk_and (mk_not t1) (mk_not t2)) :: lit))))
]
| Kapp(And,[t1;t2]) ->
let d1 = decidability t1 in
@@ -1744,7 +1746,7 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_and,
[| t1; t2; d1; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_or (mk_not t1) (mk_not t2))::lit))))
+ (loop (LocalAssum (i,mk_or (mk_not t1) (mk_not t2)) :: lit))))
]
| Kapp(Iff,[t1;t2]) ->
let d1 = decidability t1 in
@@ -1754,9 +1756,8 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_iff,
[| t1; t2; d1; d2; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,
- mk_or (mk_and t1 (mk_not t2))
- (mk_and (mk_not t1) t2))::lit))))
+ (loop (LocalAssum (i, mk_or (mk_and t1 (mk_not t2))
+ (mk_and (mk_not t1) t2)) :: lit))))
]
| Kimp(t1,t2) ->
(* t2 must be in Prop otherwise ~(t1->t2) wouldn't be ok.
@@ -1767,14 +1768,14 @@ let destructure_hyps =
[mkApp (Lazy.force coq_not_imp,
[| t1; t2; d1; mkVar i |])]);
(onClearedName i (fun i ->
- (loop ((i,None,mk_and t1 (mk_not t2)) :: lit))))
+ (loop (LocalAssum (i,mk_and t1 (mk_not t2)) :: lit))))
]
| Kapp(Not,[t]) ->
let d = decidability t in
Tacticals.New.tclTHENLIST [
Proofview.V82.tactic (generalize_tac
[mkApp (Lazy.force coq_not_not, [| t; d; mkVar i |])]);
- (onClearedName i (fun i -> (loop ((i,None,t)::lit))))
+ (onClearedName i (fun i -> (loop (LocalAssum (i,t) :: lit))))
]
| Kapp(op,[t1;t2]) ->
(try
@@ -1807,15 +1808,13 @@ let destructure_hyps =
match destructurate_type (pf_nf typ) with
| Kapp(Nat,_) ->
(Tacticals.New.tclTHEN
- (convert_hyp_no_check
- (i,body,
- (mkApp (Lazy.force coq_neq, [| t1;t2|]))))
+ (convert_hyp_no_check (set_type (mkApp (Lazy.force coq_neq, [| t1;t2|]))
+ decl))
(loop lit))
| Kapp(Z,_) ->
(Tacticals.New.tclTHEN
- (convert_hyp_no_check
- (i,body,
- (mkApp (Lazy.force coq_Zne, [| t1;t2|]))))
+ (convert_hyp_no_check (set_type (mkApp (Lazy.force coq_Zne, [| t1;t2|]))
+ decl))
(loop lit))
| _ -> loop lit
end