author | haftmann |
Fri, 27 Jun 2025 08:09:26 +0200 | |
changeset 82775 | 61c39a9e5415 |
parent 82643 | f1c14af17591 |
permissions | -rw-r--r-- |
76183 | 1 |
(* Title: HOL/Tools/SMT/lethe_replay_methods.ML |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
2 |
Author: Mathias Fleury, MPII, JKU, University Freiburg |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
3 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
4 |
Proof method for replaying veriT proofs. |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
5 |
*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
6 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
7 |
signature LETHE_REPLAY_METHODS = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
8 |
sig |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
9 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
10 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
11 |
datatype verit_rule = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
12 |
False | True | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
13 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
14 |
(*input: a repeated (normalized) assumption of assumption of in a subproof*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
15 |
Normalized_Input | Local_Input | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
16 |
(*Subproof:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
17 |
Subproof | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
18 |
(*Conjunction:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
19 |
And | Not_And | And_Pos | And_Neg | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
20 |
(*Disjunction""*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
21 |
Or | Or_Pos | Not_Or | Or_Neg | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
22 |
(*Disjunction:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
23 |
Implies | Implies_Neg1 | Implies_Neg2 | Implies_Pos | Not_Implies1 | Not_Implies2 | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
24 |
(*Equivalence:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
25 |
Equiv_neg1 | Equiv_pos1 | Equiv_pos2 | Equiv_neg2 | Not_Equiv1 | Not_Equiv2 | Equiv1 | Equiv2 | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
26 |
(*If-then-else:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
27 |
ITE_Pos1 | ITE_Pos2 | ITE_Neg1 | ITE_Neg2 | Not_ITE1 | Not_ITE2 | ITE_Intro | ITE1 | ITE2 | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
28 |
(*Equality:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
29 |
Eq_Congruent | Eq_Reflexive | Eq_Transitive | Eq_Congruent_Pred | Trans | Refl | Cong | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
30 |
(*Arithmetics:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
31 |
LA_Disequality | LA_Generic | LA_Tautology | LIA_Generic | LA_Totality | LA_RW_Eq | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
32 |
NLA_Generic | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
33 |
(*Quantifiers:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
34 |
Forall_Inst | Qnt_Rm_Unused | Qnt_Join | Onepoint | Bind | Skolem_Forall | Skolem_Ex | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
35 |
(*Resolution:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
36 |
Theory_Resolution | Resolution | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
37 |
(*Temporary rules, that the verit developers want to remove:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
38 |
AC_Simp | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
39 |
Bfun_Elim | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
40 |
Qnt_CNF | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
41 |
(*Used to introduce skolem constants*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
42 |
Definition | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
43 |
(*Former cong rules*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
44 |
Bool_Simplify | Or_Simplify | Not_Simplify | And_Simplify | Equiv_Simplify | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
45 |
Implies_Simplify | Connective_Def | Minus_Simplify | Comp_Simplify | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
46 |
Eq_Simplify | ITE_Simplify | Sum_Simplify | Prod_Simplify | All_Simplify | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
47 |
Qnt_Simplify | Not_Not | Tautological_Clause | Duplicate_Literals | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
48 |
(*Unsupported rule*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
49 |
Unsupported | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
50 |
(*For compression*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
51 |
Theory_Resolution2 | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
52 |
(*Extended rules*) |
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
53 |
Symm | Not_Symm | Reordering | Tmp_Quantifier_Simplify | |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
54 |
(*CVC5 to support new rules declared later (like BV rules)*) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
55 |
Other_Rule of string | |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
56 |
Hole |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
57 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
58 |
val requires_subproof_assms : string list -> string -> bool |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
59 |
val requires_local_input: string list -> string -> bool |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
60 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
61 |
val string_of_verit_rule: verit_rule -> string |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
62 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
63 |
type lethe_tac_args = Proof.context -> thm list -> term list -> term -> thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
64 |
type lethe_tac = Proof.context -> thm list -> term -> thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
65 |
type lethe_tac2 = Proof.context -> thm list -> term list -> term Symtab.table -> term -> thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
66 |
val bind: lethe_tac_args |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
67 |
val and_rule: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
68 |
val and_neg_rule: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
69 |
val and_pos: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
70 |
val rewrite_and_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
71 |
val rewrite_bool_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
72 |
val rewrite_comp_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
73 |
val rewrite_minus_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
74 |
val rewrite_not_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
75 |
val rewrite_eq_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
76 |
val rewrite_equiv_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
77 |
val rewrite_implies_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
78 |
val rewrite_or_simplify: lethe_tac |
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
79 |
val cong: string -> lethe_tac |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
80 |
val rewrite_connective_def: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
81 |
val duplicate_literals: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
82 |
val eq_congruent: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
83 |
val eq_congruent_pred: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
84 |
val eq_reflexive: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
85 |
val eq_transitive: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
86 |
val equiv1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
87 |
val equiv2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
88 |
val equiv_neg1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
89 |
val equiv_neg2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
90 |
val equiv_pos1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
91 |
val equiv_pos2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
92 |
val false_rule: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
93 |
val forall_inst: lethe_tac2 |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
94 |
val implies_rules: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
95 |
val implies_neg1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
96 |
val implies_neg2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
97 |
val implies_pos: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
98 |
val ite1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
99 |
val ite2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
100 |
val ite_intro: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
101 |
val ite_neg1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
102 |
val ite_neg2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
103 |
val ite_pos1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
104 |
val ite_pos2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
105 |
val rewrite_ite_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
106 |
val la_disequality: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
107 |
val la_generic: lethe_tac_args |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
108 |
val la_rw_eq: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
109 |
val lia_generic: lethe_tac |
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
110 |
val refl: string -> lethe_tac |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
111 |
val normalized_input: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
112 |
val not_and_rule: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
113 |
val not_equiv1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
114 |
val not_equiv2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
115 |
val not_implies1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
116 |
val not_implies2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
117 |
val not_ite1: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
118 |
val not_ite2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
119 |
val not_not: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
120 |
val not_or_rule: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
121 |
val or: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
122 |
val or_neg_rule: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
123 |
val or_pos_rule: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
124 |
val theory_resolution2: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
125 |
val prod_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
126 |
val qnt_join: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
127 |
val qnt_rm_unused: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
128 |
val onepoint: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
129 |
val qnt_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
130 |
val all_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
131 |
val unit_res: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
132 |
val skolem_ex: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
133 |
val skolem_forall: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
134 |
val subproof: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
135 |
val sum_simplify: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
136 |
val tautological_clause: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
137 |
val tmp_AC_rule: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
138 |
val bfun_elim: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
139 |
val qnt_cnf: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
140 |
val trans: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
141 |
val symm: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
142 |
val not_symm: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
143 |
val reordering: lethe_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
144 |
|
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
145 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
146 |
(*Extension to declare new alethe rules*) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
147 |
val declare_alethe_rule: string -> lethe_tac_args -> Context.generic -> Context.generic |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
148 |
val rm_alethe_rule: string -> Context.generic -> Context.generic |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
149 |
val get_alethe_tac: string -> Context.generic -> lethe_tac_args option |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
150 |
val print_alethe_tac: Context.generic -> Pretty.T |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
151 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
152 |
(*Useful lifting of tactics*) |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
153 |
val REPEAT_CHANGED: ('a -> tactic) -> 'a -> tactic |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
154 |
val TRY': ('a -> tactic) -> 'a -> tactic |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
155 |
|
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
156 |
val simplify_tac: Proof.context -> thm list -> int -> tactic |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
157 |
val replay_error: Proof.context -> string -> verit_rule -> thm list -> term -> 'a |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
158 |
|
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
159 |
end; |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
160 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
161 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
162 |
structure Lethe_Replay_Methods: LETHE_REPLAY_METHODS = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
163 |
struct |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
164 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
165 |
type lethe_tac_args = Proof.context -> thm list -> term list -> term -> thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
166 |
type lethe_tac = Proof.context -> thm list -> term -> thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
167 |
type lethe_tac2 = Proof.context -> thm list -> term list -> term Symtab.table -> term -> thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
168 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
169 |
(*Some general comments on the proof format: |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
170 |
1. Double negations are not always removed. This means for example that the equivalence rules |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
171 |
cannot assume that the double negations have already been removed. Therefore, we match the |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
172 |
term, instantiate the theorem, then use simp (to remove double negations), and finally use |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
173 |
assumption. |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
174 |
2. The reconstruction for rule forall_inst is buggy and tends to be very fragile, because the rule |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
175 |
is doing much more that is supposed to do. Moreover types can make trivial goals (for the |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
176 |
boolean structure) impossible to prove. |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
177 |
3. Duplicate literals are sometimes removed, mostly by the SAT solver. |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
178 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
179 |
Rules unsupported on purpose: |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
180 |
* Distinct_Elim, XOR, let (we don't need them). |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
181 |
* deep_skolemize (because it is not clear if verit still generates using it). |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
182 |
*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
183 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
184 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
185 |
datatype verit_rule = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
186 |
False | True | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
187 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
188 |
(*input: a repeated (normalized) assumption of assumption of in a subproof*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
189 |
Normalized_Input | Local_Input | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
190 |
(*Subproof:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
191 |
Subproof | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
192 |
(*Conjunction:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
193 |
And | Not_And | And_Pos | And_Neg | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
194 |
(*Disjunction""*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
195 |
Or | Or_Pos | Not_Or | Or_Neg | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
196 |
(*Disjunction:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
197 |
Implies | Implies_Neg1 | Implies_Neg2 | Implies_Pos | Not_Implies1 | Not_Implies2 | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
198 |
(*Equivalence:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
199 |
Equiv_neg1 | Equiv_pos1 | Equiv_pos2 | Equiv_neg2 | Not_Equiv1 | Not_Equiv2 | Equiv1 | Equiv2 | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
200 |
(*If-then-else:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
201 |
ITE_Pos1 | ITE_Pos2 | ITE_Neg1 | ITE_Neg2 | Not_ITE1 | Not_ITE2 | ITE_Intro | ITE1 | ITE2 | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
202 |
(*Equality:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
203 |
Eq_Congruent | Eq_Reflexive | Eq_Transitive | Eq_Congruent_Pred | Trans | Refl | Cong | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
204 |
(*Arithmetics:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
205 |
LA_Disequality | LA_Generic | LA_Tautology | LIA_Generic | LA_Totality | LA_RW_Eq | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
206 |
NLA_Generic | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
207 |
(*Quantifiers:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
208 |
Forall_Inst | Qnt_Rm_Unused | Qnt_Join | Onepoint | Bind | Skolem_Forall | Skolem_Ex | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
209 |
(*Resolution:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
210 |
Theory_Resolution | Resolution | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
211 |
(*Temporary rules, that the verit developpers want to remove:*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
212 |
AC_Simp | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
213 |
Bfun_Elim | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
214 |
Qnt_CNF | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
215 |
(*Used to introduce skolem constants*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
216 |
Definition | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
217 |
(*Former cong rules*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
218 |
Bool_Simplify | Or_Simplify | Not_Simplify | And_Simplify | Equiv_Simplify | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
219 |
Implies_Simplify | Connective_Def | Minus_Simplify | Comp_Simplify | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
220 |
Eq_Simplify | ITE_Simplify | Sum_Simplify | Prod_Simplify | All_Simplify | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
221 |
Qnt_Simplify | Not_Not | Tautological_Clause | Duplicate_Literals | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
222 |
(*Unsupported rule*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
223 |
Unsupported | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
224 |
(*For compression*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
225 |
Theory_Resolution2 | |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
226 |
(*Extended rules*) |
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
227 |
Symm | Not_Symm | Reordering | Tmp_Quantifier_Simplify | |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
228 |
(*CVC5*) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
229 |
Other_Rule of string | |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
230 |
Hole |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
231 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
232 |
fun string_of_verit_rule Bind = "Bind" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
233 |
| string_of_verit_rule Cong = "Cong" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
234 |
| string_of_verit_rule Refl = "Refl" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
235 |
| string_of_verit_rule Equiv1 = "Equiv1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
236 |
| string_of_verit_rule Equiv2 = "Equiv2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
237 |
| string_of_verit_rule Equiv_pos1 = "Equiv_pos1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
238 |
| string_of_verit_rule Equiv_pos2 = "Equiv_pos2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
239 |
| string_of_verit_rule Equiv_neg1 = "Equiv_neg1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
240 |
| string_of_verit_rule Equiv_neg2 = "Equiv_neg2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
241 |
| string_of_verit_rule Skolem_Forall = "Skolem_Forall" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
242 |
| string_of_verit_rule Skolem_Ex = "Skolem_Ex" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
243 |
| string_of_verit_rule Eq_Reflexive = "Eq_Reflexive" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
244 |
| string_of_verit_rule Theory_Resolution = "Theory_Resolution" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
245 |
| string_of_verit_rule Theory_Resolution2 = "Theory_Resolution2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
246 |
| string_of_verit_rule Forall_Inst = "forall_inst" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
247 |
| string_of_verit_rule Or = "Or" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
248 |
| string_of_verit_rule Not_Or = "Not_Or" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
249 |
| string_of_verit_rule Resolution = "Resolution" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
250 |
| string_of_verit_rule Eq_Congruent = "eq_congruent" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
251 |
| string_of_verit_rule Trans = "trans" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
252 |
| string_of_verit_rule False = "false" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
253 |
| string_of_verit_rule And = "and" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
254 |
| string_of_verit_rule And_Pos = "and_pos" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
255 |
| string_of_verit_rule Not_And = "not_and" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
256 |
| string_of_verit_rule And_Neg = "and_neg" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
257 |
| string_of_verit_rule Or_Pos = "or_pos" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
258 |
| string_of_verit_rule Or_Neg = "or_neg" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
259 |
| string_of_verit_rule AC_Simp = "ac_simp" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
260 |
| string_of_verit_rule Not_Equiv1 = "not_equiv1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
261 |
| string_of_verit_rule Not_Equiv2 = "not_equiv2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
262 |
| string_of_verit_rule Not_Implies1 = "not_implies1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
263 |
| string_of_verit_rule Not_Implies2 = "not_implies2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
264 |
| string_of_verit_rule Implies_Neg1 = "implies_neg1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
265 |
| string_of_verit_rule Implies_Neg2 = "implies_neg2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
266 |
| string_of_verit_rule Implies = "implies" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
267 |
| string_of_verit_rule Bfun_Elim = "bfun_elim" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
268 |
| string_of_verit_rule ITE1 = "ite1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
269 |
| string_of_verit_rule ITE2 = "ite2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
270 |
| string_of_verit_rule Not_ITE1 = "not_ite1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
271 |
| string_of_verit_rule Not_ITE2 = "not_ite2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
272 |
| string_of_verit_rule ITE_Pos1 = "ite_pos1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
273 |
| string_of_verit_rule ITE_Pos2 = "ite_pos2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
274 |
| string_of_verit_rule ITE_Neg1 = "ite_neg1" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
275 |
| string_of_verit_rule ITE_Neg2 = "ite_neg2" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
276 |
| string_of_verit_rule ITE_Intro = "ite_intro" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
277 |
| string_of_verit_rule LA_Disequality = "la_disequality" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
278 |
| string_of_verit_rule LA_Generic = "la_generic" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
279 |
| string_of_verit_rule LIA_Generic = "lia_generic" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
280 |
| string_of_verit_rule LA_Tautology = "la_tautology" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
281 |
| string_of_verit_rule LA_RW_Eq = "la_rw_eq" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
282 |
| string_of_verit_rule LA_Totality = "LA_Totality" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
283 |
| string_of_verit_rule NLA_Generic = "nla_generic" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
284 |
| string_of_verit_rule Eq_Transitive = "eq_transitive" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
285 |
| string_of_verit_rule Qnt_Rm_Unused = "qnt_remove_unused" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
286 |
| string_of_verit_rule Onepoint = "onepoint" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
287 |
| string_of_verit_rule Qnt_Join = "qnt_join" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
288 |
| string_of_verit_rule Eq_Congruent_Pred = "eq_congruent_pred" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
289 |
| string_of_verit_rule Normalized_Input = Lethe_Proof.normalized_input_rule |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
290 |
| string_of_verit_rule Local_Input = Lethe_Proof.local_input_rule |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
291 |
| string_of_verit_rule Subproof = "subproof" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
292 |
| string_of_verit_rule Bool_Simplify = "bool_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
293 |
| string_of_verit_rule Equiv_Simplify = "equiv_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
294 |
| string_of_verit_rule Eq_Simplify = "eq_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
295 |
| string_of_verit_rule Not_Simplify = "not_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
296 |
| string_of_verit_rule And_Simplify = "and_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
297 |
| string_of_verit_rule Or_Simplify = "or_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
298 |
| string_of_verit_rule ITE_Simplify = "ite_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
299 |
| string_of_verit_rule Implies_Simplify = "implies_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
300 |
| string_of_verit_rule Connective_Def = "connective_def" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
301 |
| string_of_verit_rule Minus_Simplify = "minus_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
302 |
| string_of_verit_rule Sum_Simplify = "sum_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
303 |
| string_of_verit_rule Prod_Simplify = "prod_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
304 |
| string_of_verit_rule All_Simplify = "all_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
305 |
| string_of_verit_rule Comp_Simplify = "comp_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
306 |
| string_of_verit_rule Qnt_Simplify = "qnt_simplify" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
307 |
| string_of_verit_rule Symm = "symm" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
308 |
| string_of_verit_rule Not_Symm = "not_symm" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
309 |
| string_of_verit_rule Reordering = "reordering" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
310 |
| string_of_verit_rule Not_Not = Lethe_Proof.not_not_rule |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
311 |
| string_of_verit_rule Tautological_Clause = "tautology" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
312 |
| string_of_verit_rule Duplicate_Literals = Lethe_Proof.contract_rule |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
313 |
| string_of_verit_rule Qnt_CNF = "qnt_cnf" |
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
314 |
| string_of_verit_rule (Other_Rule r) = r |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
315 |
| string_of_verit_rule r = "Unknown rule: " ^ \<^make_string> r |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
316 |
|
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
317 |
(** Context Extension for Rules **) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
318 |
(*We currently do not distinguish between the extension required for each solver. Maybe later |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
319 |
it will be needed.*) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
320 |
type alethe_rule_ext = {rules: (string * lethe_tac_args) list} |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
321 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
322 |
fun mk_alethe_rule_ext rules : alethe_rule_ext = |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
323 |
{rules=rules} |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
324 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
325 |
val empty_data = mk_alethe_rule_ext [] |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
326 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
327 |
fun merge_data ({rules=rules1}:alethe_rule_ext, {rules=rules2}:alethe_rule_ext) : alethe_rule_ext = |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
328 |
mk_alethe_rule_ext (AList.merge (op =) (fn _ => true) (rules1, rules2)) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
329 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
330 |
structure Data = Generic_Data |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
331 |
( |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
332 |
type T = alethe_rule_ext |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
333 |
val empty = empty_data |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
334 |
val merge = merge_data |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
335 |
) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
336 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
337 |
fun declare_alethe_rule rule tac context = |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
338 |
let |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
339 |
val {rules} = Data.get context |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
340 |
in |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
341 |
Data.map |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
342 |
(K (mk_alethe_rule_ext (AList.update (op =) (rule, tac) rules))) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
343 |
context |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
344 |
end |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
345 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
346 |
fun rm_alethe_rule stgy context = |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
347 |
let |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
348 |
val {rules} = Data.get context |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
349 |
in |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
350 |
Data.map |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
351 |
(K (mk_alethe_rule_ext (AList.delete (op =) stgy rules))) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
352 |
context |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
353 |
end |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
354 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
355 |
fun get_alethe_tac rule context = |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
356 |
let |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
357 |
val {rules} = Data.get context |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
358 |
in |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
359 |
AList.lookup (op =) rules rule |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
360 |
end |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
361 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
362 |
fun print_alethe_tac context = |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
363 |
let |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
364 |
val {rules} = Data.get context |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
365 |
in |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
366 |
rules |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
367 |
|> map (fn (a, b) => a ^ "->" ^ @{make_string} b) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
368 |
|> map Pretty.str |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
369 |
|> Pretty.big_list "Declared alethe rules:\n" |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
370 |
end |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
371 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
372 |
|
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
373 |
(** Tactics for Reconstruction **) |
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
374 |
(**) |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
375 |
fun replay_error ctxt msg rule thms t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
376 |
SMT_Replay_Methods.replay_error ctxt msg (string_of_verit_rule rule) thms t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
377 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
378 |
(* utility function *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
379 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
380 |
fun eqsubst_all ctxt thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
381 |
K ((REPEAT o HEADGOAL) (EqSubst.eqsubst_asm_tac ctxt [0] thms)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
382 |
THEN' K ((REPEAT o HEADGOAL) (EqSubst.eqsubst_tac ctxt [0] thms)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
383 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
384 |
fun simplify_tac ctxt thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
385 |
ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
386 |
|> empty_simpset |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
387 |
|> put_simpset HOL_basic_ss |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
388 |
|> (fn ctxt => ctxt addsimps @{thms simp_thms} addsimps thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
389 |
|> Simplifier.asm_full_simp_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
390 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
391 |
(* sko_forall requires the assumptions to be able to prove the equivalence in case of double |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
392 |
skolemization. See comment below. *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
393 |
fun requires_subproof_assms _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
394 |
member (op =) ["refl", "sko_forall", "sko_ex", "cong"] t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
395 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
396 |
fun requires_local_input _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
397 |
member (op =) [Lethe_Proof.local_input_rule] t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
398 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
399 |
(*This is a weaker simplification form. It is weaker, but is also less likely to loop*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
400 |
fun partial_simplify_tac ctxt thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
401 |
ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
402 |
|> empty_simpset |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
403 |
|> put_simpset HOL_basic_ss |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
404 |
|> (fn ctxt => ctxt addsimps @{thms simp_thms} addsimps thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
405 |
|> Simplifier.full_simp_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
406 |
|
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
407 |
val try_provers = SMT_Replay_Methods.try_provers |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
408 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
409 |
fun TRY' tac = fn i => TRY (tac i) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
410 |
fun REPEAT' tac = fn i => REPEAT (tac i) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
411 |
fun REPEAT_CHANGED tac = fn i => REPEAT (CHANGED (tac i)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
412 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
413 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
414 |
(* Bind *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
415 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
416 |
(*The bind rule is non-obvious due to the handling of quantifiers: |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
417 |
"\<And>x y a. x = y ==> (\<forall>b. P a b x) \<longleftrightarrow> (\<forall>b. P' a b y)" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
418 |
------------------------------------------------------ |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
419 |
\<forall>a. (\<forall>b x. P a b x) \<longleftrightarrow> (\<forall>b y. P' a b y) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
420 |
is a valid application.*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
421 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
422 |
val bind_thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
423 |
[@{lemma \<open>(\<And>x x'. x = x' \<Longrightarrow> P x = Q x') \<Longrightarrow> (\<forall>x. P x) = (\<forall>y. Q y)\<close> by blast}, |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
424 |
@{lemma \<open>(\<And>x x'. x = x' \<Longrightarrow> P x = Q x') \<Longrightarrow> (\<exists>x. P x) = (\<exists>y. Q y)\<close> by blast}, |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
425 |
@{lemma \<open>(\<And>x x'. x = x' \<Longrightarrow> P x = Q x') \<Longrightarrow> (\<exists>x. P x = Q x)\<close> by blast}, |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
426 |
@{lemma \<open>(\<And>x x'. x = x' \<Longrightarrow> P x = Q x') \<Longrightarrow> (\<forall>x. P x = Q x)\<close> by blast}] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
427 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
428 |
val bind_thms_same_name = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
429 |
[@{lemma \<open>(\<And>x. P x = Q x) \<Longrightarrow> (\<forall>x. P x) = (\<forall>y. Q y)\<close> by blast}, |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
430 |
@{lemma \<open>(\<And>x. P x = Q x) \<Longrightarrow> (\<exists>x. P x) = (\<exists>y. Q y)\<close> by blast}, |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
431 |
@{lemma \<open>(\<And>x. P x = Q x) \<Longrightarrow> (\<exists>x. P x = Q x)\<close> by blast}, |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
432 |
@{lemma \<open>(\<And>x. P x = Q x) \<Longrightarrow> (\<forall>x. P x = Q x)\<close> by blast}] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
433 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
434 |
fun extract_quantified_names_q (_ $ Abs (name, _, t)) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
435 |
apfst (curry (op ::) name) (extract_quantified_names_q t) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
436 |
| extract_quantified_names_q t = ([], t) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
437 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
438 |
fun extract_forall_quantified_names_q (Const(\<^const_name>\<open>HOL.All\<close>, _) $ Abs (name, ty, t)) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
439 |
(name, ty) :: (extract_forall_quantified_names_q t) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
440 |
| extract_forall_quantified_names_q _ = [] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
441 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
442 |
fun extract_all_forall_quantified_names_q (Const(\<^const_name>\<open>HOL.All\<close>, _) $ Abs (name, _, t)) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
443 |
name :: (extract_all_forall_quantified_names_q t) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
444 |
| extract_all_forall_quantified_names_q (t $ u) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
445 |
extract_all_forall_quantified_names_q t @ extract_all_forall_quantified_names_q u |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
446 |
| extract_all_forall_quantified_names_q _ = [] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
447 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
448 |
val extract_quantified_names_ba = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
449 |
SMT_Replay_Methods.dest_prop |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
450 |
#> extract_quantified_names_q |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
451 |
##> HOLogic.dest_eq |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
452 |
##> fst |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
453 |
##> extract_quantified_names_q |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
454 |
##> fst |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
455 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
456 |
val extract_quantified_names = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
457 |
extract_quantified_names_ba |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
458 |
#> (op @) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
459 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
460 |
val extract_all_forall_quantified_names = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
461 |
SMT_Replay_Methods.dest_prop |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
462 |
#> HOLogic.dest_eq |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
463 |
#> fst |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
464 |
#> extract_all_forall_quantified_names_q |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
465 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
466 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
467 |
fun extract_all_exists_quantified_names_q (Const(\<^const_name>\<open>HOL.Ex\<close>, _) $ Abs (name, _, t)) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
468 |
name :: (extract_all_exists_quantified_names_q t) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
469 |
| extract_all_exists_quantified_names_q (t $ u) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
470 |
extract_all_exists_quantified_names_q t @ extract_all_exists_quantified_names_q u |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
471 |
| extract_all_exists_quantified_names_q _ = [] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
472 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
473 |
val extract_all_exists_quantified_names = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
474 |
SMT_Replay_Methods.dest_prop |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
475 |
#> HOLogic.dest_eq |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
476 |
#> fst |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
477 |
#> extract_all_exists_quantified_names_q |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
478 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
479 |
|
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
480 |
fun extract_all_forall_exists_quantified_names_q (Const(\<^const_name>\<open>HOL.Ex\<close>, _) $ Abs (name, _, t)) = |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
481 |
name :: (extract_all_forall_exists_quantified_names_q t) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
482 |
| extract_all_forall_exists_quantified_names_q (Const(\<^const_name>\<open>HOL.All\<close>, _) $ Abs (name, _, t)) = |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
483 |
name :: (extract_all_forall_exists_quantified_names_q t) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
484 |
| extract_all_forall_exists_quantified_names_q (t $ u) = |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
485 |
extract_all_forall_exists_quantified_names_q t @ extract_all_forall_exists_quantified_names_q u |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
486 |
| extract_all_forall_exists_quantified_names_q _ = [] |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
487 |
|
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
488 |
val extract_bind_names = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
489 |
HOLogic.dest_eq |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
490 |
#> apply2 (fn (Free (name, _)) => name) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
491 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
492 |
fun combine_quant ctxt ((n1, n2) :: bounds) (n1' :: formula) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
493 |
TRY' (if n1 = n1' |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
494 |
then if n1 <> n2 |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
495 |
then |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
496 |
resolve_tac ctxt bind_thms |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
497 |
THEN' TRY'(resolve_tac ctxt [@{thm refl}]) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
498 |
THEN' combine_quant ctxt bounds formula |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
499 |
else resolve_tac ctxt bind_thms_same_name THEN' combine_quant ctxt bounds formula |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
500 |
else resolve_tac ctxt @{thms allI} THEN' combine_quant ctxt ((n1, n2) :: bounds) formula) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
501 |
| combine_quant _ _ _ = K all_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
502 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
503 |
fun bind_quants ctxt args t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
504 |
combine_quant ctxt (map extract_bind_names args) (extract_quantified_names t) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
505 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
506 |
fun generalize_prems_q [] prems = prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
507 |
| generalize_prems_q (_ :: quants) prems = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
508 |
generalize_prems_q quants (@{thm spec} OF [prems]) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
509 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
510 |
fun generalize_prems t = generalize_prems_q (fst (extract_quantified_names_ba t)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
511 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
512 |
fun bind ctxt [prems] args t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
513 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
514 |
bind_quants ctxt args t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
515 |
THEN' TRY' (SOLVED' (resolve_tac ctxt [generalize_prems t prems] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
516 |
THEN_ALL_NEW TRY' (assume_tac ctxt ORELSE' resolve_tac ctxt @{thms refl})))) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
517 |
| bind ctxt thms _ t = replay_error ctxt "invalid bind application" Bind thms t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
518 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
519 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
520 |
(* Congruence/Refl *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
521 |
|
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
522 |
fun cong name ctxt thms = try_provers name ctxt |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
523 |
(string_of_verit_rule Cong) [ |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
524 |
("basic", SMT_Replay_Methods.cong_basic ctxt thms), |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
525 |
("unfolding then reflexivity", SMT_Replay_Methods.cong_unfolding_trivial ctxt thms), |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
526 |
("unfolding then auto", SMT_Replay_Methods.cong_unfolding_first ctxt thms), |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
527 |
("full", SMT_Replay_Methods.cong_full ctxt thms)] thms |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
528 |
|
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
529 |
fun refl name ctxt thm t = |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
530 |
(case find_first (fn thm => t = Thm.full_prop_of thm) thm of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
531 |
SOME thm => thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
532 |
| NONE => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
533 |
(case try (fn t => SMT_Replay_Methods.match_instantiate ctxt t @{thm refl}) t of |
78177
ea7a3cc64df5
early inclusion of cvc5 proof reconstruction; slightly reorganize smt/z3_reals;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
76183
diff
changeset
|
534 |
NONE => cong name ctxt thm t |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
535 |
| SOME thm => thm)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
536 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
537 |
(* Instantiation *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
538 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
539 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
540 |
fun dropWhile _ [] = [] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
541 |
| dropWhile f (x :: xs) = if f x then dropWhile f xs else x :: xs |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
542 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
543 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
544 |
fun forall_inst ctxt _ _ insts t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
545 |
let |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
546 |
fun instantiate_and_solve i ({context = ctxt, prems = [prem], ...}: Subgoal.focus) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
547 |
let |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
548 |
val t = Thm.prop_of prem |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
549 |
val quants = t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
550 |
|> SMT_Replay_Methods.dest_prop |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
551 |
|> extract_forall_quantified_names_q |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
552 |
val insts = map (Symtab.lookup insts o fst) (rev quants) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
553 |
|> dropWhile (curry (op =) NONE) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
554 |
|> rev |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
555 |
fun option_map _ NONE = NONE |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
556 |
| option_map f (SOME a) = SOME (f a) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
557 |
fun instantiate_with inst prem = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
558 |
Drule.infer_instantiate' ctxt [NONE, inst] @{thm spec} OF [prem] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
559 |
val inst_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
560 |
fold instantiate_with |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
561 |
(map (option_map (Thm.cterm_of ctxt)) insts) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
562 |
prem |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
563 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
564 |
(Method.insert_tac ctxt [inst_thm] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
565 |
THEN' TRY' (fn i => assume_tac ctxt i) |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
566 |
THEN' TRY' (partial_simplify_tac ctxt @{thms eq_commute ac_simps}) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
567 |
THEN' TRY' (blast_tac ctxt)) i |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
568 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
569 |
| instantiate_and_solve _ ({context = ctxt, prems = thms, ...}: Subgoal.focus) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
570 |
replay_error ctxt "invalid application" Forall_Inst thms t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
571 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
572 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
573 |
match_tac ctxt [@{thm disj_not1} RSN (1, @{thm iffD2}) OF [@{thm impI}]] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
574 |
THEN' (fn i => Subgoal.FOCUS (instantiate_and_solve i) ctxt i)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
575 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
576 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
577 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
578 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
579 |
(* Or *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
580 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
581 |
fun or _ (thm :: _) _ = thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
582 |
| or ctxt thms t = replay_error ctxt "invalid bind application" Or thms t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
583 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
584 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
585 |
(* Implication *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
586 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
587 |
val implies_pos_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
588 |
[@{lemma \<open>\<not>(A \<longrightarrow> B) \<or> \<not>A \<or> B\<close> by blast}, |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
589 |
@{lemma \<open>\<not>(\<not>A \<longrightarrow> B) \<or> A \<or> B\<close> by blast}] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
590 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
591 |
fun implies_pos ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
592 |
resolve_tac ctxt implies_pos_thm) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
593 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
594 |
(* Skolemization *) |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
595 |
local |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
596 |
fun split _ [] = ([], []) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
597 |
| split f (a :: xs) = |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
598 |
split f xs |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
599 |
|> (if f a then apfst (curry (op ::) a) else apsnd (curry (op ::) a)) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
600 |
in |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
601 |
fun extract_rewrite_rule_assumption _ thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
602 |
let |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
603 |
fun is_rewrite_rule thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
604 |
(case Thm.prop_of thm of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
605 |
\<^term>\<open>Trueprop\<close> $ (Const(\<^const_name>\<open>HOL.eq\<close>, _) $ _ $ Free(_, _)) => true |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
606 |
| _ => false) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
607 |
fun is_context_rule thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
608 |
(case Thm.prop_of thm of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
609 |
\<^term>\<open>Trueprop\<close> $ (Const(\<^const_name>\<open>HOL.eq\<close>, _) $ Free(_, _) $ Free(_, _)) => true |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
610 |
| _ => false) |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
611 |
val (ctxt_eq, other) = |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
612 |
thms |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
613 |
|> split is_context_rule |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
614 |
val (rew, other) = |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
615 |
other |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
616 |
|> split is_rewrite_rule |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
617 |
in |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
618 |
(ctxt_eq, rew, other) |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
619 |
end |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
620 |
end |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
621 |
(* |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
622 |
Without compression, we have to rewrite skolems only once. However, it can happen than the same |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
623 |
skolem constant is used multiple times with a different name under the forall. |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
624 |
|
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
625 |
For strictness, we use the multiple rewriting only when compressing is activated. |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
626 |
*) |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
627 |
local |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
628 |
fun rewrite_all_skolems thm_indirect ctxt ((v,SOME thm) :: thms) = |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
629 |
let |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
630 |
val rewrite_sk_thms = |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
631 |
List.mapPartial (fn tm => SOME (tm OF [thm]) handle THM _ => NONE) thm_indirect |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
632 |
val multiple_rew = if SMT_Config.compress_verit_proofs ctxt then REPEAT_CHANGED else fn x => x |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
633 |
in |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
634 |
multiple_rew (EqSubst.eqsubst_tac ctxt [0] rewrite_sk_thms |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
635 |
THEN' SOLVED' (K (HEADGOAL (partial_simplify_tac ctxt (@{thms eq_commute}))))) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
636 |
THEN' rewrite_all_skolems thm_indirect ctxt thms |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
637 |
end |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
638 |
| rewrite_all_skolems thm_indirect ctxt ((_,NONE) :: thms) = rewrite_all_skolems thm_indirect ctxt thms |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
639 |
| rewrite_all_skolems _ _ [] = K (all_tac) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
640 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
641 |
fun extract_var_name (thm :: thms) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
642 |
let val name = Thm.concl_of thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
643 |
|> SMT_Replay_Methods.dest_prop |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
644 |
|> HOLogic.dest_eq |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
645 |
|> fst |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
646 |
|> (fn Const(_,_) $ Abs(name, _, _) => [(name,@{thm sym} OF [thm])] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
647 |
| _ => []) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
648 |
in name :: extract_var_name thms end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
649 |
| extract_var_name [] = [] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
650 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
651 |
fun skolem_tac extractor thm1 thm2 ctxt thms t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
652 |
let |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
653 |
val (ctxt_eq, ts, other) = extract_rewrite_rule_assumption ctxt thms |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
654 |
|
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
655 |
fun ordered_definitions () = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
656 |
let |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
657 |
val var_order = extractor t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
658 |
val thm_names_with_var = extract_var_name ts |> flat |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
659 |
in map (fn v => (v, AList.lookup (op =) thm_names_with_var v)) var_order end |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
660 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
661 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
662 |
K (unfold_tac ctxt ctxt_eq) |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
663 |
THEN' rewrite_all_skolems thm2 ctxt (ordered_definitions ()) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
664 |
THEN' (eqsubst_all ctxt (map (fn thm => thm RS sym) other)) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
665 |
THEN_ALL_NEW TRY' (resolve_tac ctxt @{thms refl}) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
666 |
THEN' K (unfold_tac ctxt ctxt_eq) |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
667 |
THEN' TRY' (partial_simplify_tac ctxt (@{thms eq_commute}))) |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
668 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
669 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
670 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
671 |
val skolem_forall = |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
672 |
skolem_tac extract_all_forall_exists_quantified_names_q @{thm verit_sko_forall_indirect} |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
673 |
@{thms verit_sko_forall_indirect2 verit_sko_ex_indirect2} |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
674 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
675 |
val skolem_ex = |
75956
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
676 |
skolem_tac extract_all_forall_exists_quantified_names_q @{thm verit_sko_ex_indirect} |
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
75299
diff
changeset
|
677 |
@{thms verit_sko_ex_indirect2 verit_sko_forall_indirect2} |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
678 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
679 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
680 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
681 |
fun eq_reflexive ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t @{thm refl} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
682 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
683 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
684 |
fun not_not_prove ctxt _ = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
685 |
K (unfold_tac ctxt @{thms not_not}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
686 |
THEN' match_tac ctxt @{thms verit_or_simplify_1} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
687 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
688 |
fun duplicate_literals_prove ctxt prems = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
689 |
Method.insert_tac ctxt prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
690 |
THEN' match_tac ctxt @{thms ccontr} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
691 |
THEN' K (unfold_tac ctxt @{thms de_Morgan_disj not_not}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
692 |
THEN' TRY o CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt @{thms conjE}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
693 |
THEN' CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt @{thms disjE}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
694 |
THEN' REPEAT' (ematch_tac ctxt @{thms notE} THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
695 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
696 |
fun tautological_clause_prove ctxt _ = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
697 |
match_tac ctxt @{thms verit_or_neg} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
698 |
THEN' K (unfold_tac ctxt @{thms not_not disj_assoc[symmetric]}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
699 |
THEN' TRY' (match_tac ctxt @{thms notE} THEN_ALL_NEW assume_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
700 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
701 |
val theory_resolution2_lemma = @{lemma \<open>a \<Longrightarrow> a = b \<Longrightarrow> b\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
702 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
703 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
704 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
705 |
fun prove_abstract abstracter tac ctxt thms t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
706 |
let |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
707 |
val thms = map (Conv.fconv_rule Thm.eta_long_conversion) thms |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
708 |
val t' = Thm.eta_long_conversion (Object_Logic.dest_judgment ctxt (Thm.cterm_of ctxt t)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
709 |
val (_, t2) = Logic.dest_equals (Thm.prop_of t') |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
710 |
val thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
711 |
SMT_Replay_Methods.prove_abstract ctxt thms t2 tac ( |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
712 |
fold_map (abstracter o SMT_Replay_Methods.dest_thm) thms ##>> |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
713 |
abstracter (SMT_Replay_Methods.dest_prop t2)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
714 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
715 |
@{thm verit_Pure_trans} OF [t', thm] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
716 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
717 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
718 |
val not_not = prove_abstract SMT_Replay_Methods.abstract_bool_shallow not_not_prove |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
719 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
720 |
val tautological_clause = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
721 |
prove_abstract SMT_Replay_Methods.abstract_bool_shallow tautological_clause_prove |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
722 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
723 |
val duplicate_literals = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
724 |
prove_abstract SMT_Replay_Methods.abstract_bool_shallow duplicate_literals_prove |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
725 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
726 |
val unit_res = prove_abstract SMT_Replay_Methods.abstract_bool_shallow SMT_Replay_Methods.prop_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
727 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
728 |
(*match_instantiate does not work*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
729 |
fun theory_resolution2 ctxt prems t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
730 |
SMT_Replay_Methods.prove ctxt t (fn _ => match_tac ctxt [theory_resolution2_lemma OF prems]) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
731 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
732 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
733 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
734 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
735 |
fun normalized_input ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
736 |
Method.insert_tac ctxt prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
737 |
THEN' TRY' (K (unfold_tac ctxt @{thms SMT.trigger_def})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
738 |
THEN' TRY' (partial_simplify_tac ctxt @{thms eq_commute})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
739 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
740 |
val false_rule_thm = @{lemma \<open>\<not>False\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
741 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
742 |
fun false_rule ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t false_rule_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
743 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
744 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
745 |
(* Transitivity *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
746 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
747 |
val trans_bool_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
748 |
@{lemma \<open>P = Q \<Longrightarrow> Q \<Longrightarrow> P\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
749 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
750 |
fun trans ctxt thms t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
751 |
let |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
752 |
val prop_of = HOLogic.dest_Trueprop o Thm.full_prop_of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
753 |
fun combine_thms [thm1, thm2] = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
754 |
(case (prop_of thm1, prop_of thm2) of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
755 |
((Const(\<^const_name>\<open>HOL.eq\<close>, _) $ t1 $ t2), |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
756 |
(Const(\<^const_name>\<open>HOL.eq\<close>, _) $ t3 $ t4)) => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
757 |
if t2 aconv t3 then thm1 RSN (1, thm2 RSN (2, @{thm trans})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
758 |
else if t2 aconv t4 then thm1 RSN ((1, ((thm2 RS sym)) RSN (2, @{thm trans}))) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
759 |
else if t1 aconv t4 then thm2 RSN (1, thm1 RSN (2, @{thm trans})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
760 |
else raise Fail "invalid trans theorem" |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
761 |
| _ => trans_bool_thm OF [thm1, thm2]) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
762 |
| combine_thms (thm1 :: thm2 :: thms) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
763 |
combine_thms (combine_thms [thm1, thm2] :: thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
764 |
| combine_thms thms = replay_error ctxt "invalid bind application" Trans thms t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
765 |
val t' = Thm.eta_long_conversion (Object_Logic.dest_judgment ctxt (Thm.cterm_of ctxt t)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
766 |
val (_, t2) = Logic.dest_equals (Thm.prop_of t') |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
767 |
val thms = map (Conv.fconv_rule Thm.eta_long_conversion) thms |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
768 |
val trans_thm = combine_thms thms |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
769 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
770 |
(case (prop_of trans_thm, t2) of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
771 |
((Const(\<^const_name>\<open>HOL.eq\<close>, _) $ t1 $ _), |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
772 |
(Const(\<^const_name>\<open>HOL.eq\<close>, _) $ t3 $ _)) => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
773 |
if t1 aconv t3 then trans_thm else trans_thm RS sym |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
774 |
| _ => trans_thm (*to be on the safe side*)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
775 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
776 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
777 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
778 |
fun tmp_AC_rule ctxt thms t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
779 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
780 |
Method.insert_tac ctxt thms |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
781 |
THEN' REPEAT_ALL_NEW (partial_simplify_tac ctxt @{thms eq_commute} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
782 |
THEN' TRY' (simplify_tac ctxt @{thms ac_simps conj_ac}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
783 |
THEN' TRY' (Classical.fast_tac ctxt))) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
784 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
785 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
786 |
(* And/Or *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
787 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
788 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
789 |
fun not_and_rule_prove ctxt prems = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
790 |
Method.insert_tac ctxt prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
791 |
THEN' K (unfold_tac ctxt @{thms de_Morgan_conj}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
792 |
THEN_ALL_NEW TRY' (assume_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
793 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
794 |
fun or_pos_prove ctxt _ = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
795 |
K (unfold_tac ctxt @{thms de_Morgan_disj not_not}) |
79576
157de27b0863
fix reconstruction of Alethe's and_pos rule
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
78177
diff
changeset
|
796 |
THEN' match_tac ctxt @{thms verit_and_pos verit_farkas} |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
797 |
THEN' K (unfold_tac ctxt @{thms de_Morgan_conj de_Morgan_disj not_not}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
798 |
THEN' TRY' (assume_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
799 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
800 |
fun not_or_rule_prove ctxt prems = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
801 |
Method.insert_tac ctxt prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
802 |
THEN' K (unfold_tac ctxt @{thms de_Morgan_disj not_not}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
803 |
THEN' TRY' (REPEAT' (match_tac ctxt @{thms conjI})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
804 |
THEN_ALL_NEW ((REPEAT' (ematch_tac ctxt @{thms conjE})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
805 |
THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
806 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
807 |
fun and_rule_prove ctxt prems = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
808 |
Method.insert_tac ctxt prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
809 |
THEN' (fn i => REPEAT (dresolve_tac ctxt @{thms conjE} i THEN assume_tac ctxt (i+1))) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
810 |
THEN' TRY' (assume_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
811 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
812 |
fun and_neg_rule_prove ctxt _ = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
813 |
match_tac ctxt @{thms verit_and_neg} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
814 |
THEN' K (unfold_tac ctxt @{thms de_Morgan_conj not_not}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
815 |
THEN' TRY' (assume_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
816 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
817 |
fun prover tac = prove_abstract SMT_Replay_Methods.abstract_prop tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
818 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
819 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
820 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
821 |
val and_rule = prover and_rule_prove |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
822 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
823 |
val not_and_rule = prover not_and_rule_prove |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
824 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
825 |
val not_or_rule = prover not_or_rule_prove |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
826 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
827 |
val or_pos_rule = prover or_pos_prove |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
828 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
829 |
val and_neg_rule = prover and_neg_rule_prove |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
830 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
831 |
val or_neg_rule = prove_abstract SMT_Replay_Methods.abstract_bool (fn ctxt => fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
832 |
resolve_tac ctxt @{thms verit_or_neg} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
833 |
THEN' K (unfold_tac ctxt @{thms not_not}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
834 |
THEN_ALL_NEW |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
835 |
(REPEAT' |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
836 |
(SOLVED' (match_tac ctxt @{thms disjI1} THEN_ALL_NEW assume_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
837 |
ORELSE' (match_tac ctxt @{thms disjI2} THEN' TRY' (assume_tac ctxt))))) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
838 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
839 |
fun and_pos ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
840 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
841 |
REPEAT_CHANGED (resolve_tac ctxt @{thms verit_and_pos}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
842 |
THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
843 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
844 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
845 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
846 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
847 |
(* Equivalence Transformation *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
848 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
849 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
850 |
fun prove_equiv equiv_thm ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
851 |
Method.insert_tac ctxt [equiv_thm OF [thm]] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
852 |
THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
853 |
| prove_equiv _ ctxt thms t = replay_error ctxt "invalid application in some equiv rule" Unsupported thms t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
854 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
855 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
856 |
val not_equiv1 = prove_equiv @{lemma \<open>\<not>(A \<longleftrightarrow> B) \<Longrightarrow> A \<or> B\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
857 |
val not_equiv2 = prove_equiv @{lemma \<open>\<not>(A \<longleftrightarrow> B) \<Longrightarrow> \<not>A \<or> \<not>B\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
858 |
val equiv1 = prove_equiv @{lemma \<open>(A \<longleftrightarrow> B) \<Longrightarrow> \<not>A \<or> B\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
859 |
val equiv2 = prove_equiv @{lemma \<open>(A \<longleftrightarrow> B) \<Longrightarrow> A \<or> \<not>B\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
860 |
val not_implies1 = prove_equiv @{lemma \<open>\<not>(A \<longrightarrow> B) \<Longrightarrow> A\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
861 |
val not_implies2 = prove_equiv @{lemma \<open>\<not>(A \<longrightarrow>B) \<Longrightarrow> \<not>B\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
862 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
863 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
864 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
865 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
866 |
(* Equivalence Lemma *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
867 |
(*equiv_pos2 is very important for performance. We have tried various tactics, including |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
868 |
a specialisation of SMT_Replay_Methods.match_instantiate, but there never was a measurable |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
869 |
and consistent gain.*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
870 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
871 |
fun prove_equiv_pos_neg2 thm ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
872 |
SMT_Replay_Methods.match_instantiate ctxt t thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
873 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
874 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
875 |
val equiv_pos1_thm = @{lemma \<open>\<not>(a \<longleftrightarrow> b) \<or> a \<or> \<not>b\<close> by blast+} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
876 |
val equiv_pos1 = prove_equiv_pos_neg2 equiv_pos1_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
877 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
878 |
val equiv_pos2_thm = @{lemma \<open>\<And>a b. (a \<noteq> b) \<or> \<not>a \<or> b\<close> by blast+} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
879 |
val equiv_pos2 = prove_equiv_pos_neg2 equiv_pos2_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
880 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
881 |
val equiv_neg1_thm = @{lemma \<open>(a \<longleftrightarrow> b) \<or> \<not>a \<or> \<not>b\<close> by blast+} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
882 |
val equiv_neg1 = prove_equiv_pos_neg2 equiv_neg1_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
883 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
884 |
val equiv_neg2_thm = @{lemma \<open>(a \<longleftrightarrow> b) \<or> a \<or> b\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
885 |
val equiv_neg2 = prove_equiv_pos_neg2 equiv_neg2_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
886 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
887 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
888 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
889 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
890 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
891 |
fun implies_pos_neg_term ctxt thm (\<^term>\<open>Trueprop\<close> $ |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
892 |
(\<^term>\<open>HOL.disj\<close> $ (\<^term>\<open>HOL.implies\<close> $ a $ b) $ _)) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
893 |
Drule.infer_instantiate' ctxt (map (SOME o Thm.cterm_of ctxt) [a, b]) thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
894 |
| implies_pos_neg_term ctxt thm t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
895 |
replay_error ctxt "invalid application in Implies" Unsupported [thm] t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
896 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
897 |
fun prove_implies_pos_neg thm ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
898 |
let val thm = implies_pos_neg_term ctxt thm t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
899 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
900 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
901 |
Method.insert_tac ctxt [thm] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
902 |
THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
903 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
904 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
905 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
906 |
val implies_neg1_thm = @{lemma \<open>(a \<longrightarrow> b) \<or> a\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
907 |
val implies_neg1 = prove_implies_pos_neg implies_neg1_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
908 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
909 |
val implies_neg2_thm = @{lemma \<open>(a \<longrightarrow> b) \<or> \<not>b\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
910 |
val implies_neg2 = prove_implies_pos_neg implies_neg2_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
911 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
912 |
val implies_thm = @{lemma \<open>(a \<longrightarrow> b) \<Longrightarrow> \<not>a \<or> b\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
913 |
fun implies_rules ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
914 |
Method.insert_tac ctxt [implies_thm OF prems] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
915 |
THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
916 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
917 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
918 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
919 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
920 |
(* BFun *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
921 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
922 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
923 |
val bfun_theorems = @{thms verit_bfun_elim} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
924 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
925 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
926 |
fun bfun_elim ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
927 |
Method.insert_tac ctxt prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
928 |
THEN' TRY' (eqsubst_all ctxt bfun_theorems) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
929 |
THEN' TRY' (simplify_tac ctxt @{thms eq_commute all_conj_distrib ex_disj_distrib})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
930 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
931 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
932 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
933 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
934 |
(* If-Then-Else *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
935 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
936 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
937 |
fun prove_ite ite_thm thm ctxt = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
938 |
resolve_tac ctxt [ite_thm OF [thm]] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
939 |
THEN' TRY' (assume_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
940 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
941 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
942 |
val ite_pos1_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
943 |
@{lemma \<open>\<not>(if x then P else Q) \<or> x \<or> Q\<close> by auto} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
944 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
945 |
fun ite_pos1 ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
946 |
resolve_tac ctxt [ite_pos1_thm]) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
947 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
948 |
val ite_pos2_thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
949 |
@{lemma \<open>\<not>(if x then P else Q) \<or> \<not>x \<or> P\<close> \<open>\<not>(if \<not>x then P else Q) \<or> x \<or> P\<close> by auto} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
950 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
951 |
fun ite_pos2 ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => resolve_tac ctxt ite_pos2_thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
952 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
953 |
val ite_neg1_thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
954 |
@{lemma \<open>(if x then P else Q) \<or> x \<or> \<not>Q\<close> \<open>(if x then P else \<not>Q) \<or> x \<or> Q\<close> by auto} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
955 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
956 |
fun ite_neg1 ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => resolve_tac ctxt ite_neg1_thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
957 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
958 |
val ite_neg2_thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
959 |
@{lemma \<open>(if x then P else Q) \<or> \<not>x \<or> \<not>P\<close> \<open>(if \<not>x then P else Q) \<or> x \<or> \<not>P\<close> |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
960 |
\<open>(if x then \<not>P else Q) \<or> \<not>x \<or> P\<close> \<open>(if \<not>x then \<not>P else Q) \<or> x \<or> P\<close> |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
961 |
by auto} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
962 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
963 |
fun ite_neg2 ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => resolve_tac ctxt ite_neg2_thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
964 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
965 |
val ite1_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
966 |
@{lemma \<open>(if x then P else Q) \<Longrightarrow> x \<or> Q\<close> by (auto split: if_splits) } |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
967 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
968 |
fun ite1 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => prove_ite ite1_thm thm ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
969 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
970 |
val ite2_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
971 |
@{lemma \<open>(if x then P else Q) \<Longrightarrow> \<not>x \<or> P\<close> by (auto split: if_splits) } |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
972 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
973 |
fun ite2 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => prove_ite ite2_thm thm ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
974 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
975 |
val not_ite1_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
976 |
@{lemma \<open>\<not>(if x then P else Q) \<Longrightarrow> x \<or> \<not>Q\<close> by (auto split: if_splits) } |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
977 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
978 |
fun not_ite1 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => prove_ite not_ite1_thm thm ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
979 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
980 |
val not_ite2_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
981 |
@{lemma \<open>\<not>(if x then P else Q) \<Longrightarrow> \<not>x \<or> \<not>P\<close> by (auto split: if_splits) } |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
982 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
983 |
fun not_ite2 ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => prove_ite not_ite2_thm thm ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
984 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
985 |
(*ite_intro can introduce new terms that are in the proof exactly the sames as the old one, but are |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
986 |
not internally, hence the possible reordering.*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
987 |
fun ite_intro ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
988 |
let |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
989 |
fun simplify_ite ctxt thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
990 |
ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
991 |
|> empty_simpset |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
992 |
|> put_simpset HOL_basic_ss |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
993 |
|> (fn ctxt => ctxt addsimps thms @ @{thms if_True if_False refl simp_thms if_cancel}) |
82643
f1c14af17591
prefer Simplifier over bootstrap-only Raw_Simplifier
haftmann
parents:
80701
diff
changeset
|
994 |
|> Simplifier.add_eqcong @{thm verit_ite_if_cong} |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
995 |
|> Simplifier.full_simp_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
996 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
997 |
SMT_Replay_Methods.prove ctxt t (fn _ => simplify_ite ctxt [] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
998 |
THEN' TRY' (simplify_ite ctxt @{thms eq_commute})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
999 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1000 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1001 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1002 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1003 |
(* Quantifiers *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1004 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1005 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1006 |
val rm_unused = @{lemma \<open>(\<forall>x. P) = P\<close> \<open>(\<exists>x. P) = P\<close> by blast+} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1007 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1008 |
fun qnt_cnf_tac ctxt = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1009 |
simplify_tac ctxt @{thms de_Morgan_conj de_Morgan_disj imp_conv_disj |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1010 |
iff_conv_conj_imp if_bool_eq_disj ex_simps all_simps ex_disj_distrib |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1011 |
verit_connective_def(3) all_conj_distrib} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1012 |
THEN' TRY' (Blast.blast_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1013 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1014 |
fun qnt_rm_unused ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1015 |
K (unfold_tac ctxt rm_unused)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1016 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1017 |
fun onepoint ctxt prems t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1018 |
Method.insert_tac ctxt prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1019 |
THEN' Simplifier.full_simp_tac (put_simpset HOL_basic_ss (empty_simpset ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1020 |
addsimps @{thms HOL.simp_thms HOL.all_simps} |
80701 | 1021 |
|> Simplifier.add_proc @{simproc HOL.defined_All} |
1022 |
|> Simplifier.add_proc @{simproc HOL.defined_Ex}) |
|
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1023 |
THEN' TRY' (Blast.blast_tac ctxt) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1024 |
THEN' TRY' (Metis_Tactic.metis_tac [] ATP_Problem_Generate.combsN ctxt [])) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1025 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1026 |
fun qnt_join ctxt _ t = SMT_Replay_Methods.prove ctxt t Classical.fast_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1027 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1028 |
fun qnt_cnf ctxt _ t = SMT_Replay_Methods.prove ctxt t qnt_cnf_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1029 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1030 |
fun qnt_simplify ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1031 |
partial_simplify_tac ctxt []) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1032 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1033 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1034 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1035 |
(* Equality *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1036 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1037 |
fun eq_transitive ctxt _ t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1038 |
REPEAT_CHANGED (resolve_tac ctxt [@{thm disj_not1} RSN (1, @{thm iffD2}) OF @{thms impI}]) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1039 |
THEN' REPEAT' (resolve_tac ctxt @{thms impI}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1040 |
THEN' REPEAT' (eresolve_tac ctxt @{thms conjI}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1041 |
THEN' REPEAT' (fn i => dresolve_tac ctxt @{thms verit_eq_transitive} i THEN assume_tac ctxt (i+1)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1042 |
THEN' resolve_tac ctxt @{thms refl}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1043 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1044 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1045 |
fun find_rew rews t t' = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1046 |
(case AList.lookup (op =) rews (t, t') of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1047 |
SOME thm => SOME (thm COMP @{thm symmetric}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1048 |
| NONE => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1049 |
(case AList.lookup (op =) rews (t', t) of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1050 |
SOME thm => SOME thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1051 |
| NONE => NONE)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1052 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1053 |
fun eq_pred_conv rews t ctxt ctrm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1054 |
(case find_rew rews t (Thm.term_of ctrm) of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1055 |
SOME thm => Conv.rewr_conv thm ctrm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1056 |
| NONE => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1057 |
(case t of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1058 |
f $ arg => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1059 |
(Conv.fun_conv (eq_pred_conv rews f ctxt) then_conv |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1060 |
Conv.arg_conv (eq_pred_conv rews arg ctxt)) ctrm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1061 |
| Abs (_, _, f) => Conv.abs_conv (eq_pred_conv rews f o snd) ctxt ctrm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1062 |
| _ => Conv.all_conv ctrm)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1063 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1064 |
fun eq_pred_rewrite_tac ({context = ctxt, prems, concl, ...}: Subgoal.focus) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1065 |
let |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1066 |
val rews = prems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1067 |
|> map_filter (try (apfst (HOLogic.dest_eq o Thm.term_of o Object_Logic.dest_judgment ctxt o |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1068 |
Thm.cconcl_of) o `(fn x => x))) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1069 |
|> map (apsnd (fn x => @{thm eq_reflection} OF [x])) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1070 |
fun conv_left conv = Conv.arg_conv (Conv.arg_conv conv) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1071 |
fun conv_left_neg conv = Conv.arg_conv (Conv.arg_conv (Conv.arg_conv conv)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1072 |
val (t1, conv_term) = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1073 |
(case Thm.term_of (Object_Logic.dest_judgment ctxt concl) of |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1074 |
Const(_, _) $ (Const(\<^const_name>\<open>HOL.Not\<close>, _) $ t1) $ _ => (t1, conv_left) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1075 |
| Const(_, _) $ t1 $ (Const(\<^const_name>\<open>HOL.Not\<close>, _) $ _) => (t1, conv_left_neg) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1076 |
| Const(_, _) $ t1 $ _ => (t1, conv_left) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1077 |
| t1 => (t1, conv_left)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1078 |
fun throwing_CONVERSION cv i st = Seq.single (Conv.gconv_rule cv i st) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1079 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1080 |
throwing_CONVERSION (conv_term (eq_pred_conv rews t1 ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1081 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1082 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1083 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1084 |
fun eq_congruent_pred ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1085 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1086 |
REPEAT' (resolve_tac ctxt [@{thm disj_not1[of \<open>_ = _\<close>]} RSN (1, @{thm iffD2}) OF @{thms impI}]) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1087 |
THEN' (fn i => Subgoal.FOCUS (fn focus => eq_pred_rewrite_tac focus i) ctxt i) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1088 |
THEN' (resolve_tac ctxt @{thms refl excluded_middle excluded_middle[of \<open>\<not>_\<close>, unfolded not_not]} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1089 |
ORELSE' assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1090 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1091 |
val eq_congruent = eq_congruent_pred |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1092 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1093 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1094 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1095 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1096 |
(* Subproof *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1097 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1098 |
fun subproof ctxt [prem] t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1099 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1100 |
(REPEAT' (resolve_tac ctxt [@{thm disj_not1} RSN (1, @{thm iffD2}) OF [@{thm impI}], |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1101 |
@{thm disj_not1[of \<open>\<not>_\<close>, unfolded not_not]} RSN (1, @{thm iffD2}) OF [@{thm impI}]]) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1102 |
THEN' resolve_tac ctxt [prem] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1103 |
THEN_ALL_NEW assume_tac ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1104 |
THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1105 |
ORELSE' TRY' (Method.insert_tac ctxt [prem] THEN' Blast.blast_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1106 |
| subproof ctxt prems t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1107 |
replay_error ctxt "invalid application, only one assumption expected" Subproof prems t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1108 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1109 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1110 |
(* Simplifying Steps *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1111 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1112 |
(* The reconstruction is a bit tricky: At first we only rewrite only based on the rules that are |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1113 |
passed as argument. Then we simply use simp_thms to reconstruct all the proofs (and these theorems |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1114 |
cover all the simplification below). |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1115 |
*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1116 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1117 |
fun rewrite_only_thms ctxt thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1118 |
ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1119 |
|> empty_simpset |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1120 |
|> put_simpset HOL_basic_ss |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1121 |
|> (fn ctxt => ctxt addsimps thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1122 |
|> Simplifier.full_simp_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1123 |
fun rewrite_only_thms_tmp ctxt thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1124 |
rewrite_only_thms ctxt thms |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1125 |
THEN' TRY' (K (Clasimp.auto_tac ctxt)) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1126 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1127 |
fun rewrite_thms ctxt thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1128 |
ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1129 |
|> empty_simpset |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1130 |
|> put_simpset HOL_basic_ss |
82643
f1c14af17591
prefer Simplifier over bootstrap-only Raw_Simplifier
haftmann
parents:
80701
diff
changeset
|
1131 |
|> Simplifier.add_eqcong @{thm eq_reflection[OF imp_cong]} |
75299
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1132 |
|> (fn ctxt => ctxt addsimps thms addsimps @{thms simp_thms}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1133 |
|> Simplifier.full_simp_tac |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1134 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1135 |
fun chain_rewrite_thms ctxt thms = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1136 |
TRY' (rewrite_only_thms ctxt thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1137 |
THEN' TRY' (rewrite_thms ctxt thms) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1138 |
THEN' TRY' (K (Clasimp.auto_tac ctxt)) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1139 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1140 |
fun chain_rewrite_two_step_with_ac_simps ctxt thms1 thms2 = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1141 |
TRY' (rewrite_only_thms ctxt thms1) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1142 |
THEN' TRY' (rewrite_thms ctxt thms2) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1143 |
THEN' TRY' (K (Clasimp.auto_tac ctxt)) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1144 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1145 |
fun chain_rewrite_thms_two_step ctxt thms1 thms2 thms3 thms4 = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1146 |
chain_rewrite_two_step_with_ac_simps ctxt thms1 thms2 |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1147 |
THEN' TRY' (chain_rewrite_two_step_with_ac_simps ctxt thms3 thms4) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1148 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1149 |
val imp_refl = @{lemma \<open>(P \<longrightarrow> P) = True\<close> by blast} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1150 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1151 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1152 |
fun rewrite_bool_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1153 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1154 |
(chain_rewrite_thms ctxt @{thms verit_bool_simplify})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1155 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1156 |
fun rewrite_and_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1157 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1158 |
(chain_rewrite_two_step_with_ac_simps ctxt @{thms verit_and_simplify verit_and_simplify1} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1159 |
@{thms verit_and_simplify})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1160 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1161 |
fun rewrite_or_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1162 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1163 |
(chain_rewrite_two_step_with_ac_simps ctxt @{thms verit_or_simplify verit_or_simplify_1} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1164 |
@{thms verit_or_simplify}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1165 |
THEN' TRY' (REPEAT' (match_tac ctxt @{thms verit_and_pos}) THEN' assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1166 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1167 |
fun rewrite_not_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1168 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1169 |
(rewrite_only_thms_tmp ctxt @{thms verit_not_simplify})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1170 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1171 |
fun rewrite_equiv_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1172 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1173 |
(rewrite_only_thms_tmp ctxt @{thms verit_equiv_simplify})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1174 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1175 |
fun rewrite_eq_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1176 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1177 |
(chain_rewrite_two_step_with_ac_simps ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1178 |
@{thms verit_eq_simplify} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1179 |
(Named_Theorems.get ctxt @{named_theorems ac_simps}))) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1180 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1181 |
fun rewrite_implies_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1182 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1183 |
(rewrite_only_thms_tmp ctxt @{thms verit_implies_simplify})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1184 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1185 |
(* It is more efficient to first fold the implication symbols. |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1186 |
That is however not enough when symbols appears within |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1187 |
expressions, hence we also unfold implications in the |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1188 |
other direction. *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1189 |
fun rewrite_connective_def ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1190 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1191 |
chain_rewrite_thms_two_step ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1192 |
(imp_refl :: @{thms not_not verit_connective_def[symmetric]}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1193 |
(@{thms verit_connective_def[symmetric]}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1194 |
(imp_refl :: @{thms not_not verit_connective_def}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1195 |
(@{thms verit_connective_def})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1196 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1197 |
fun rewrite_minus_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1198 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1199 |
chain_rewrite_two_step_with_ac_simps ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1200 |
@{thms arith_simps verit_minus_simplify} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1201 |
(Named_Theorems.get ctxt @{named_theorems ac_simps} @ |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1202 |
@{thms numerals arith_simps arith_special |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1203 |
numeral_class.numeral.numeral_One})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1204 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1205 |
fun rewrite_comp_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1206 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1207 |
chain_rewrite_thms ctxt @{thms verit_comp_simplify}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1208 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1209 |
fun rewrite_ite_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1210 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1211 |
(rewrite_only_thms_tmp ctxt @{thms verit_ite_simplify})) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1212 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1213 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1214 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1215 |
(* Simplifications *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1216 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1217 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1218 |
fun simplify_arith ctxt thms = partial_simplify_tac ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1219 |
(thms @ Named_Theorems.get ctxt @{named_theorems ac_simps} @ @{thms arith_simps}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1220 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1221 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1222 |
fun sum_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1223 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1224 |
simplify_arith ctxt @{thms verit_sum_simplify arith_special} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1225 |
THEN' TRY' (K (Clasimp.auto_tac ctxt))) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1226 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1227 |
fun prod_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1228 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1229 |
simplify_arith ctxt @{thms verit_prod_simplify} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1230 |
THEN' TRY' (K (Clasimp.auto_tac ctxt))) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1231 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1232 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1233 |
fun all_simplify ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1234 |
SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1235 |
TRY' (K (Clasimp.auto_tac ctxt))) (*TODO: TEMP until cvc5 produces fine grained simplification steps *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1236 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1237 |
(* Arithmetics *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1238 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1239 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1240 |
val la_rw_eq_thm = @{lemma \<open>(a :: nat) = b \<or> (a \<le> b) \<or> (a \<ge> b)\<close> by auto} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1241 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1242 |
fun la_rw_eq ctxt _ t = SMT_Replay_Methods.match_instantiate ctxt t la_rw_eq_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1243 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1244 |
fun la_generic_prove args ctxt _ = SMT_Replay_Arith.la_farkas args ctxt |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1245 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1246 |
fun la_generic ctxt _ args = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1247 |
prove_abstract (SMT_Replay_Methods.abstract_arith_shallow ctxt) (la_generic_prove args) ctxt [] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1248 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1249 |
fun lia_generic ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1250 |
SMT_Replay_Methods.prove_arith_rewrite SMT_Replay_Methods.abstract_neq ctxt t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1251 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1252 |
fun la_disequality ctxt _ t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1253 |
SMT_Replay_Methods.match_instantiate ctxt t @{thm verit_la_disequality} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1254 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1255 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1256 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1257 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1258 |
(* Symm and Not_Symm*) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1259 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1260 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1261 |
fun prove_symm symm_thm ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ => |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1262 |
Method.insert_tac ctxt [symm_thm OF [thm]] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1263 |
THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1264 |
| prove_symm _ ctxt thms t = replay_error ctxt "invalid application in some symm rule" Unsupported thms t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1265 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1266 |
val symm_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1267 |
@{lemma \<open>(B = A) \<Longrightarrow> A = B\<close> by blast } |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1268 |
val symm = prove_symm symm_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1269 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1270 |
val not_symm_thm = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1271 |
@{lemma \<open>\<not>(B = A) \<Longrightarrow> \<not>(A = B)\<close> by blast } |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1272 |
val not_symm = prove_symm not_symm_thm |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1273 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1274 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1275 |
(* Reordering *) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1276 |
local |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1277 |
fun prove_reordering ctxt [thm] t = SMT_Replay_Methods.prove ctxt t (fn _ =>( |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1278 |
Method.insert_tac ctxt [thm] |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1279 |
THEN' match_tac ctxt @{thms ccontr} |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1280 |
THEN' K (unfold_tac ctxt @{thms de_Morgan_disj}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1281 |
THEN' TRY o CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt @{thms conjE}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1282 |
THEN' CHANGED_PROP o REPEAT_ALL_NEW (ematch_tac ctxt @{thms disjE}) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1283 |
THEN' REPEAT'(ematch_tac ctxt @{thms notE} THEN' TRY' (assume_tac ctxt)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1284 |
)) |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1285 |
| prove_reordering ctxt thms t = |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1286 |
replay_error ctxt "invalid application in some reordering rule" Unsupported thms t |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1287 |
in |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1288 |
val reordering = prove_reordering |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1289 |
end |
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1290 |
|
da591621d6ae
split veriT reconstruction into Lethe and veriT part
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
diff
changeset
|
1291 |
end; |