| author | wenzelm | 
| Mon, 31 Dec 2018 12:02:31 +0100 | |
| changeset 69556 | 0a38f23ca4c5 | 
| parent 69205 | 8050734eee3e | 
| child 69593 | 3dda49e08b9d | 
| permissions | -rw-r--r-- | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Tools/SMT/conj_disj_perm.ML  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
2  | 
Author: Sascha Boehme, TU Muenchen  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
3  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
4  | 
Tactic to prove equivalence of permutations of conjunctions and disjunctions.  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
7  | 
signature CONJ_DISJ_PERM =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
8  | 
sig  | 
| 60752 | 9  | 
val conj_disj_perm_tac: Proof.context -> int -> tactic  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
10  | 
end  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
12  | 
structure Conj_Disj_Perm: CONJ_DISJ_PERM =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
13  | 
struct  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
14  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
15  | 
fun with_assumption ct f =  | 
| 67149 | 16  | 
let val ct' = Thm.apply \<^cterm>\<open>HOL.Trueprop\<close> ct  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
17  | 
in Thm.implies_intr ct' (f (Thm.assume ct')) end  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
18  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
19  | 
fun eq_from_impls thm1 thm2 = thm2 INCR_COMP (thm1 INCR_COMP @{thm iffI})
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
20  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
21  | 
fun add_lit thm = Termtab.update (HOLogic.dest_Trueprop (Thm.prop_of thm), thm)  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
22  | 
|
| 67091 | 23  | 
val ndisj1_rule = @{lemma "\<not>(P \<or> Q) \<Longrightarrow> \<not>P" by auto}
 | 
24  | 
val ndisj2_rule = @{lemma "\<not>(P \<or> Q) \<Longrightarrow> \<not>Q" by auto}
 | 
|
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
26  | 
fun explode_thm thm =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
27  | 
(case HOLogic.dest_Trueprop (Thm.prop_of thm) of  | 
| 67149 | 28  | 
    \<^const>\<open>HOL.conj\<close> $ _ $ _ => explode_conj_thm @{thm conjunct1} @{thm conjunct2} thm
 | 
29  | 
| \<^const>\<open>HOL.Not\<close> $ (\<^const>\<open>HOL.disj\<close> $ _ $ _) => explode_conj_thm ndisj1_rule ndisj2_rule thm  | 
|
30  | 
  | \<^const>\<open>HOL.Not\<close> $ (\<^const>\<open>HOL.Not\<close> $ _) => explode_thm (thm RS @{thm notnotD})
 | 
|
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
31  | 
| _ => add_lit thm)  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
32  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
33  | 
and explode_conj_thm rule1 rule2 thm lits =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
34  | 
explode_thm (thm RS rule1) (explode_thm (thm RS rule2) (add_lit thm lits))  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
35  | 
|
| 67091 | 36  | 
val not_false_rule = @{lemma "\<not>False" by auto}
 | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
37  | 
fun explode thm = explode_thm thm (add_lit not_false_rule (add_lit @{thm TrueI} Termtab.empty))
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
38  | 
|
| 67149 | 39  | 
fun find_dual_lit lits (\<^const>\<open>HOL.Not\<close> $ t, thm) = Termtab.lookup lits t |> Option.map (pair thm)  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
40  | 
| find_dual_lit _ _ = NONE  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
41  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
42  | 
fun find_dual_lits lits = Termtab.get_first (find_dual_lit lits) lits  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
43  | 
|
| 67091 | 44  | 
val not_not_rule = @{lemma "P \<Longrightarrow> \<not>\<not>P" by auto}
 | 
45  | 
val ndisj_rule = @{lemma "\<not>P \<Longrightarrow> \<not>Q \<Longrightarrow> \<not>(P \<or> Q)" by auto}
 | 
|
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
46  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
47  | 
fun join lits t =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
48  | 
(case Termtab.lookup lits t of  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
49  | 
SOME thm => thm  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
50  | 
| NONE => join_term lits t)  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
51  | 
|
| 67149 | 52  | 
and join_term lits (\<^const>\<open>HOL.conj\<close> $ t $ u) = @{thm conjI} OF (map (join lits) [t, u])
 | 
53  | 
| join_term lits (\<^const>\<open>HOL.Not\<close> $ (\<^const>\<open>HOL.disj\<close> $ t $ u)) =  | 
|
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
54  | 
ndisj_rule OF (map (join lits o HOLogic.mk_not) [t, u])  | 
| 67149 | 55  | 
| join_term lits (\<^const>\<open>HOL.Not\<close> $ (\<^const>\<open>HOL.Not\<close> $ t)) = join lits t RS not_not_rule  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
56  | 
  | join_term _ t = raise TERM ("join_term", [t])
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
57  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
58  | 
fun prove_conj_disj_imp ct cu = with_assumption ct (fn thm => join (explode thm) (Thm.term_of cu))  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
59  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
60  | 
fun prove_conj_disj_eq (clhs, crhs) =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
61  | 
let  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
62  | 
val thm1 = prove_conj_disj_imp clhs crhs  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
63  | 
val thm2 = prove_conj_disj_imp crhs clhs  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
64  | 
in eq_from_impls thm1 thm2 end  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
65  | 
|
| 67091 | 66  | 
val not_not_false_rule = @{lemma "\<not>\<not>False \<Longrightarrow> P" by auto}
 | 
67  | 
val not_true_rule = @{lemma "\<not>True \<Longrightarrow> P" by auto}
 | 
|
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
68  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
69  | 
fun prove_any_imp ct =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
70  | 
(case Thm.term_of ct of  | 
| 67149 | 71  | 
    \<^const>\<open>HOL.False\<close> => @{thm FalseE}
 | 
72  | 
| \<^const>\<open>HOL.Not\<close> $ (\<^const>\<open>HOL.Not\<close> $ \<^const>\<open>HOL.False\<close>) => not_not_false_rule  | 
|
73  | 
| \<^const>\<open>HOL.Not\<close> $ \<^const>\<open>HOL.True\<close> => not_true_rule  | 
|
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
74  | 
  | _ => raise CTERM ("prove_any_imp", [ct]))
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
75  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
76  | 
fun prove_contradiction_imp ct =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
77  | 
with_assumption ct (fn thm =>  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
78  | 
let val lits = explode thm  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
79  | 
in  | 
| 67149 | 80  | 
(case Termtab.lookup lits \<^const>\<open>HOL.False\<close> of  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
81  | 
        SOME thm' => thm' RS @{thm FalseE}
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
82  | 
| NONE =>  | 
| 67149 | 83  | 
(case Termtab.lookup lits (\<^const>\<open>HOL.Not\<close> $ \<^const>\<open>HOL.True\<close>) of  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
84  | 
SOME thm' => thm' RS not_true_rule  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
85  | 
| NONE =>  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
86  | 
(case find_dual_lits lits of  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
87  | 
                SOME (not_lit_thm, lit_thm) => @{thm notE} OF [not_lit_thm, lit_thm]
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
88  | 
              | NONE => raise CTERM ("prove_contradiction", [ct]))))
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
89  | 
end)  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
90  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
91  | 
fun prove_contradiction_eq to_right (clhs, crhs) =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
92  | 
let  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
93  | 
val thm1 = if to_right then prove_contradiction_imp clhs else prove_any_imp clhs  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
94  | 
val thm2 = if to_right then prove_any_imp crhs else prove_contradiction_imp crhs  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
95  | 
in eq_from_impls thm1 thm2 end  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
96  | 
|
| 67091 | 97  | 
val contrapos_rule = @{lemma "(\<not>P) = (\<not>Q) \<Longrightarrow> P = Q" by auto}
 | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
98  | 
fun contrapos prove cp = contrapos_rule OF [prove (apply2 (Thm.apply @{cterm HOL.Not}) cp)]
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
100  | 
datatype kind = True | False | Conj | Disj | Other  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
101  | 
|
| 67149 | 102  | 
fun choose t _ _ _ \<^const>\<open>HOL.True\<close> = t  | 
103  | 
| choose _ f _ _ \<^const>\<open>HOL.False\<close> = f  | 
|
104  | 
| choose _ _ c _ (\<^const>\<open>HOL.conj\<close> $ _ $ _) = c  | 
|
105  | 
| choose _ _ _ d (\<^const>\<open>HOL.disj\<close> $ _ $ _) = d  | 
|
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
106  | 
| choose _ _ _ _ _ = Other  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
107  | 
|
| 67149 | 108  | 
fun kind_of (\<^const>\<open>HOL.Not\<close> $ t) = choose False True Disj Conj t  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
109  | 
| kind_of t = choose True False Conj Disj t  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
110  | 
|
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
111  | 
fun prove_conj_disj_perm ct cp =  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
112  | 
(case apply2 (kind_of o Thm.term_of) cp of  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
113  | 
(Conj, Conj) => prove_conj_disj_eq cp  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
114  | 
| (Disj, Disj) => contrapos prove_conj_disj_eq cp  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
115  | 
| (Conj, False) => prove_contradiction_eq true cp  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
116  | 
| (False, Conj) => prove_contradiction_eq false cp  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
117  | 
| (Disj, True) => contrapos (prove_contradiction_eq true) cp  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
118  | 
| (True, Disj) => contrapos (prove_contradiction_eq false) cp  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
119  | 
  | _ => raise CTERM ("prove_conj_disj_perm", [ct]))
 | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
120  | 
|
| 60752 | 121  | 
fun conj_disj_perm_tac ctxt = CSUBGOAL (fn (ct, i) =>  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
122  | 
(case Thm.term_of ct of  | 
| 67149 | 123  | 
    \<^const>\<open>HOL.Trueprop\<close> $ (@{const HOL.eq(bool)} $ _ $ _) =>
 | 
| 60752 | 124  | 
resolve_tac ctxt [prove_conj_disj_perm ct (Thm.dest_binop (Thm.dest_arg ct))] i  | 
| 
59381
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
125  | 
| _ => no_tac))  | 
| 
 
de4218223e00
more efficient Z3 proof reconstruction for rewrite steps with conjunctions and disjunctions, similar to Old_SMT
 
boehmes 
parents:  
diff
changeset
 | 
126  | 
|
| 
69205
 
8050734eee3e
add reconstruction by veriT in method smt
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
67149 
diff
changeset
 | 
127  | 
end;  |