| author | desharna | 
| Mon, 20 Mar 2023 15:02:17 +0100 | |
| changeset 77697 | f35cbb4da88a | 
| parent 76183 | 8089593a364a | 
| child 78177 | ea7a3cc64df5 | 
| 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*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 53 | Symm | Not_Symm | Reordering | Tmp_Quantifier_Simplify | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 54 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 55 | 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 | 56 | 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 | 57 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 58 | 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 | 59 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 60 | 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 | 61 | 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 | 62 | 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 | 63 | 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 | 64 | 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 | 65 | 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 | 66 | 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 | 67 | 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 | 68 | 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 | 69 | 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 | 70 | 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 | 71 | 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 | 72 | 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 | 73 | 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 | 74 | 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 | 75 | val rewrite_or_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 cong: 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_connective_def: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 78 | 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 | 79 | 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 | 80 | 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 | 81 | 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 | 82 | 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 | 83 | val equiv1: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 84 | val equiv2: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 85 | 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 | 86 | 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 | 87 | 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 | 88 | 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 | 89 | 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 | 90 | 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 | 91 | 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 | 92 | 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 | 93 | 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 | 94 | 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 | 95 | val ite1: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 96 | val ite2: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 97 | 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 | 98 | 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 | 99 | 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 | 100 | 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 | 101 | 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 | 102 | 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 | 103 | 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 | 104 | 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 | 105 | 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 | 106 | val lia_generic: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 107 | val refl: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 108 | 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 | 109 | 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 | 110 | 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 | 111 | 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 | 112 | 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 | 113 | 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 | 114 | 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 | 115 | 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 | 116 | 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 | 117 | 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 | 118 | val or: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 119 | 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 | 120 | 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 | 121 | 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 | 122 | 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 | 123 | 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 | 124 | 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 | 125 | val onepoint: 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_simplify: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 127 | 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 | 128 | 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 | 129 | 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 | 130 | 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 | 131 | val subproof: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 132 | 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 | 133 | 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 | 134 | 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 | 135 | 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 | 136 | 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 | 137 | val trans: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 138 | val symm: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 139 | 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 | 140 | val reordering: lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 141 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 142 | (* | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 143 | val : lethe_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 144 | *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 145 |   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 | 146 |   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 | 147 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 148 | end; | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 149 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 150 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 151 | 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 | 152 | struct | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 153 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 154 | 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 | 155 | 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 | 156 | 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 | 157 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 158 | (*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 | 159 | 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 | 160 | 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 | 161 | 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 | 162 | assumption. | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 163 | 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 | 164 | 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 | 165 | 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 | 166 | 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 | 167 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 168 | Rules unsupported on purpose: | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 169 | * 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 | 170 | * 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 | 171 | *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 172 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 173 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 174 | datatype verit_rule = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 175 | False | True | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 176 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 177 | (*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 | 178 | Normalized_Input | Local_Input | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 179 | (*Subproof:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 180 | Subproof | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 181 | (*Conjunction:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 182 | 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 | 183 | (*Disjunction""*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 184 | 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 | 185 | (*Disjunction:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 186 | 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 | 187 | (*Equivalence:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 188 | 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 | 189 | (*If-then-else:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 190 | 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 | 191 | (*Equality:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 192 | 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 | 193 | (*Arithmetics:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 194 | 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 | 195 | NLA_Generic | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 196 | (*Quantifiers:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 197 | 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 | 198 | (*Resolution:*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 199 | Theory_Resolution | Resolution | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 200 | (*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 | 201 | AC_Simp | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 202 | Bfun_Elim | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 203 | Qnt_CNF | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 204 | (*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 | 205 | Definition | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 206 | (*Former cong rules*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 207 | 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 | 208 | 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 | 209 | 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 | 210 | 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 | 211 | (*Unsupported rule*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 212 | Unsupported | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 213 | (*For compression*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 214 | Theory_Resolution2 | | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 215 | (*Extended rules*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 216 | Symm | Not_Symm | Reordering | Tmp_Quantifier_Simplify | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 217 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 218 | 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 | 219 | | 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 | 220 | | 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 | 221 | | 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 | 222 | | 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 | 223 | | 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 | 224 | | 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 | 225 | | 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 | 226 | | 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 | 227 | | 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 | 228 | | 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 | 229 | | 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 | 230 | | 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 | 231 | | 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 | 232 | | 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 | 233 | | 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 | 234 | | 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 | 235 | | 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 | 236 | | 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 | 237 | | 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 | 238 | | 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 | 239 | | 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 | 240 | | 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 | 241 | | 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 | 242 | | 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 | 243 | | 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 | 244 | | 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 | 245 | | 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 | 246 | | 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 | 247 | | 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 | 248 | | 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 | 249 | | 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 | 250 | | 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 | 251 | | 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 | 252 | | 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 | 253 | | 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 | 254 | | 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 | 255 | | 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 | 256 | | 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 | 257 | | 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 | 258 | | 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 | 259 | | 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 | 260 | | 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 | 261 | | 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 | 262 | | 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 | 263 | | 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 | 264 | | 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 | 265 | | 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 | 266 | | 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 | 267 | | 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 | 268 | | 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 | 269 | | 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 | 270 | | 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 | 271 | | 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 | 272 | | 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 | 273 | | 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 | 274 | | 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 | 275 | | 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 | 276 | | 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 | 277 | | 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 | 278 | | 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 | 279 | | 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 | 280 | | 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 | 281 | | 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 | 282 | | 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 | 283 | | 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 | 284 | | 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 | 285 | | 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 | 286 | | 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 | 287 | | 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 | 288 | | 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 | 289 | | 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 | 290 | | 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 | 291 | | 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 | 292 | | 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 | 293 | | 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 | 294 | | 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 | 295 | | 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 | 296 | | 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 | 297 | | 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 | 298 | | 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 | 299 | | string_of_verit_rule Qnt_CNF = "qnt_cnf" | 
| 
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 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 | 301 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 302 | 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 | 303 | 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 | 304 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 305 | (* utility function *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 306 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 307 | 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 | 308 | 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 | 309 | 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 | 310 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 311 | 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 | 312 | ctxt | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 313 | |> empty_simpset | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 314 | |> 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 | 315 |   |> (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 | 316 | |> 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 | 317 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 318 | (* 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 | 319 | skolemization. See comment below. *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 320 | 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 | 321 | 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 | 322 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 323 | 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 | 324 | 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 | 325 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 326 | (*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 | 327 | 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 | 328 | ctxt | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 329 | |> empty_simpset | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 330 | |> 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 | 331 |   |> (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 | 332 | |> Simplifier.full_simp_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 333 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 334 | val try_provers = SMT_Replay_Methods.try_provers "verit" | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 335 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 336 | 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 | 337 | 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 | 338 | 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 | 339 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 340 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 341 | (* Bind *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 342 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 343 | (*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 | 344 | "\<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 | 345 | ------------------------------------------------------ | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 346 | \<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 | 347 | is a valid application.*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 348 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 349 | val bind_thms = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 350 |   [@{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 | 351 |    @{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 | 352 |    @{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 | 353 |    @{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 | 354 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 355 | 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 | 356 |   [@{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 | 357 |    @{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 | 358 |    @{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 | 359 |    @{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 | 360 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 361 | 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 | 362 | 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 | 363 | | 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 | 364 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 365 | 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 | 366 | (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 | 367 | | 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 | 368 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 369 | 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 | 370 | 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 | 371 | | 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 | 372 | 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 | 373 | | 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 | 374 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 375 | 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 | 376 | 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 | 377 | #> extract_quantified_names_q | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 378 | ##> HOLogic.dest_eq | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 379 | ##> fst | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 380 | ##> extract_quantified_names_q | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 381 | ##> fst | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 382 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 383 | val extract_quantified_names = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 384 | extract_quantified_names_ba | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 385 | #> (op @) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 386 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 387 | 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 | 388 | 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 | 389 | #> HOLogic.dest_eq | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 390 | #> fst | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 391 | #> 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 | 392 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 393 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 394 | 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 | 395 | 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 | 396 | | 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 | 397 | 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 | 398 | | 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 | 399 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 400 | 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 | 401 | 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 | 402 | #> HOLogic.dest_eq | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 403 | #> fst | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 404 | #> 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 | 405 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 406 | |
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 407 | 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: 
75299diff
changeset | 408 | 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: 
75299diff
changeset | 409 | | 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: 
75299diff
changeset | 410 | 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: 
75299diff
changeset | 411 | | 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: 
75299diff
changeset | 412 | 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: 
75299diff
changeset | 413 | | extract_all_forall_exists_quantified_names_q _ = [] | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 414 | |
| 75299 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 415 | val extract_bind_names = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 416 | HOLogic.dest_eq | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 417 | #> 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 | 418 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 419 | 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 | 420 | TRY' (if n1 = n1' | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 421 | then if n1 <> n2 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 422 | then | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 423 | 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 | 424 |          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 | 425 | 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 | 426 | 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 | 427 |      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 | 428 | | 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 | 429 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 430 | 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 | 431 | 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 | 432 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 433 | 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 | 434 | | 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 | 435 |       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 | 436 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 437 | 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 | 438 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 439 | 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 | 440 | 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 | 441 | 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 | 442 | 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 | 443 |       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 | 444 | | 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 | 445 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 446 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 447 | (* Congruence/Refl *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 448 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 449 | fun cong ctxt thms = try_provers ctxt | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 450 | (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 | 451 |   ("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 | 452 |   ("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 | 453 |   ("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 | 454 |   ("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 | 455 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 456 | fun refl ctxt thm t = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 457 | (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 | 458 | SOME thm => thm | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 459 | | NONE => | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 460 |         (case try (fn t => SMT_Replay_Methods.match_instantiate ctxt t @{thm refl}) t of
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 461 | NONE => cong ctxt thm t | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 462 | | SOME thm => thm)) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 463 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 464 | (* Instantiation *) | 
| 
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 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 467 | fun dropWhile _ [] = [] | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 468 | | 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 | 469 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 470 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 471 | 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 | 472 | let | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 473 |     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 | 474 | let | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 475 | 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 | 476 | val quants = t | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 477 | |> 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 | 478 | |> 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 | 479 | 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 | 480 | |> dropWhile (curry (op =) NONE) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 481 | |> rev | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 482 | 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 | 483 | | 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 | 484 | 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 | 485 |             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 | 486 | val inst_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 487 | fold instantiate_with | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 488 | (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 | 489 | prem | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 490 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 491 | (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 | 492 | 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: 
75299diff
changeset | 493 |            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: 
75299diff
changeset | 494 | 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 | 495 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 496 |      | 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 | 497 | 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 | 498 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 499 | 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 | 500 |       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 | 501 | 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 | 502 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 503 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 504 | |
| 
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 | (* Or *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 507 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 508 | fun or _ (thm :: _) _ = thm | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 509 | | 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 | 510 | |
| 
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 | (* Implication *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 513 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 514 | val implies_pos_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 515 |   [@{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 | 516 |   @{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 | 517 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 518 | 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 | 519 | 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 | 520 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 521 | (* Skolemization *) | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 522 | local | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 523 | fun split _ [] = ([], []) | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 524 | | split f (a :: xs) = | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 525 | split f xs | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 526 | |> (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: 
75299diff
changeset | 527 | in | 
| 75299 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 528 | 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 | 529 | let | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 530 | 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 | 531 | (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 | 532 | \<^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 | 533 | | _ => false) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 534 | 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 | 535 | (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 | 536 | \<^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 | 537 | | _ => false) | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 538 | 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 | 539 | thms | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 540 | |> split is_context_rule | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 541 | val (rew, other) = | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 542 | other | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 543 | |> 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 | 544 | in | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 545 | (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 | 546 | end | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 547 | end | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 548 | (* | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 549 | 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: 
75299diff
changeset | 550 | 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 | 551 | |
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 552 | 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: 
75299diff
changeset | 553 | *) | 
| 75299 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 554 | local | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 555 | 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: 
75299diff
changeset | 556 | let | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 557 | val rewrite_sk_thms = | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 558 | 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: 
75299diff
changeset | 559 | 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: 
75299diff
changeset | 560 | in | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 561 | 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: 
75299diff
changeset | 562 |          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: 
75299diff
changeset | 563 | 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: 
75299diff
changeset | 564 | end | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 565 | | 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 | 566 | | 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 | 567 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 568 | 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 | 569 | 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 | 570 | |> 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 | 571 | |> HOLogic.dest_eq | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 572 | |> fst | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 573 |          |> (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 | 574 | | _ => []) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 575 | 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 | 576 | | extract_var_name [] = [] | 
| 
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 | 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 | 579 | let | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 580 | 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: 
75299diff
changeset | 581 | |
| 75299 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 582 | fun ordered_definitions () = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 583 | let | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 584 | 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 | 585 | 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: 
75299diff
changeset | 586 | 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 | 587 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 588 | 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 | 589 | K (unfold_tac ctxt ctxt_eq) | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 590 | 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: 
75299diff
changeset | 591 | 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: 
75299diff
changeset | 592 |       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: 
75299diff
changeset | 593 | THEN' K (unfold_tac ctxt ctxt_eq) | 
| 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 594 |       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 | 595 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 596 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 597 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 598 | val skolem_forall = | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 599 |   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: 
75299diff
changeset | 600 |     @{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 | 601 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 602 | val skolem_ex = | 
| 75956 
1e2a9d2251b0
remove duplicate parsing for alethe; fix skolemization;
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
75299diff
changeset | 603 |   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: 
75299diff
changeset | 604 |     @{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 | 605 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 606 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 607 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 608 | 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 | 609 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 610 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 611 | 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 | 612 |     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 | 613 |     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 | 614 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 615 | 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 | 616 | 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 | 617 |     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 | 618 |     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 | 619 |     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 | 620 |     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 | 621 |     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 | 622 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 623 | 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 | 624 |     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 | 625 |     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 | 626 |     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 | 627 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 628 |   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 | 629 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 630 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 631 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 632 | 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 | 633 | let | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 634 | 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 | 635 | 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 | 636 | 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 | 637 | val thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 638 | 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 | 639 | 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 | 640 | 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 | 641 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 642 |     @{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 | 643 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 644 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 645 | 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 | 646 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 647 | val tautological_clause = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 648 | 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 | 649 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 650 | val duplicate_literals = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 651 | 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 | 652 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 653 | 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 | 654 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 655 | (*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 | 656 | 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 | 657 | 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 | 658 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 659 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 660 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 661 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 662 | 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 | 663 | 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 | 664 |   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 | 665 |   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 | 666 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 667 | 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 | 668 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 669 | 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 | 670 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 671 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 672 | (* Transitivity *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 673 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 674 | val trans_bool_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 675 |   @{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 | 676 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 677 | 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 | 678 | let | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 679 | 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 | 680 | 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 | 681 | (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 | 682 | ((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 | 683 | (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 | 684 |             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 | 685 |             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 | 686 |             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 | 687 | 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 | 688 | | _ => 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 | 689 | | 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 | 690 | 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 | 691 | | 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 | 692 | 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 | 693 | 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 | 694 | 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 | 695 | 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 | 696 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 697 | (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 | 698 | ((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 | 699 | (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 | 700 | 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 | 701 | | _ => 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 | 702 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 703 | |
| 
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 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 | 706 | 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 | 707 | 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 | 708 |     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 | 709 |     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 | 710 | 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 | 711 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 712 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 713 | (* And/Or *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 714 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 715 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 716 | 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 | 717 | 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 | 718 |      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 | 719 | 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 | 720 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 721 | 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 | 722 |      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 | 723 |      THEN' match_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 | 724 |      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 | 725 | 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 | 726 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 727 | 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 | 728 | 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 | 729 |      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 | 730 |      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 | 731 |      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 | 732 | 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 | 733 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 734 | 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 | 735 | 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 | 736 |      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 | 737 | 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 | 738 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 739 | 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 | 740 |      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 | 741 |      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 | 742 | 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 | 743 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 744 | 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 | 745 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 746 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 747 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 748 | 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 | 749 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 750 | 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 | 751 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 752 | 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 | 753 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 754 | 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 | 755 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 756 | 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 | 757 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 758 | 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 | 759 |   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 | 760 |   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 | 761 | THEN_ALL_NEW | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 762 | (REPEAT' | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 763 |       (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 | 764 |        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 | 765 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 766 | 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 | 767 | 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 | 768 |   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 | 769 | 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 | 770 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 771 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 772 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 773 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 774 | (* Equivalence Transformation *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 775 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 776 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 777 | 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 | 778 | 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 | 779 | 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 | 780 | | 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 | 781 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 782 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 783 | 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 | 784 | 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 | 785 | 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 | 786 | 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 | 787 | 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 | 788 | 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 | 789 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 790 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 791 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 792 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 793 | (* Equivalence Lemma *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 794 | (*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 | 795 | 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 | 796 | and consistent gain.*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 797 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 798 | 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 | 799 | 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 | 800 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 801 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 802 | 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 | 803 | 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 | 804 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 805 | 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 | 806 | 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 | 807 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 808 | 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 | 809 | 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 | 810 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 811 | 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 | 812 | 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 | 813 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 814 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 815 | |
| 
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 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 818 | 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 | 819 | (\<^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 | 820 | 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 | 821 | | 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 | 822 | 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 | 823 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 824 | 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 | 825 | 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 | 826 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 827 | 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 | 828 | 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 | 829 | 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 | 830 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 831 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 832 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 833 | 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 | 834 | 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 | 835 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 836 | 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 | 837 | 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 | 838 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 839 | 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 | 840 | 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 | 841 | 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 | 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 | (* BFun *) | 
| 
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 |   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 | 851 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 852 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 853 | 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 | 854 | 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 | 855 | 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 | 856 |   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 | 857 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 858 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 859 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 860 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 861 | (* If-Then-Else *) | 
| 
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 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 864 | 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 | 865 | 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 | 866 | 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 | 867 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 868 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 869 | val ite_pos1_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 870 |   @{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 | 871 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 872 | 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 | 873 | 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 | 874 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 875 | val ite_pos2_thms = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 876 |   @{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 | 877 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 878 | 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 | 879 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 880 | val ite_neg1_thms = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 881 |   @{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 | 882 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 883 | 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 | 884 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 885 | val ite_neg2_thms = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 886 |   @{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 | 887 | \<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 | 888 | by auto} | 
| 
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 | 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 | 891 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 892 | val ite1_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 893 |   @{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 | 894 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 895 | 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 | 896 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 897 | val ite2_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 898 |   @{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 | 899 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 900 | 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 | 901 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 902 | val not_ite1_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 903 |   @{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 | 904 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 905 | 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 | 906 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 907 | val not_ite2_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 908 |   @{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 | 909 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 910 | 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 | 911 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 912 | (*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 | 913 | 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 | 914 | 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 | 915 | let | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 916 | 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 | 917 | ctxt | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 918 | |> empty_simpset | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 919 | |> 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 | 920 |       |> (fn ctxt => ctxt addsimps thms @ @{thms if_True if_False refl simp_thms if_cancel})
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 921 |       |> Raw_Simplifier.add_eqcong @{thm verit_ite_if_cong}
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 922 | |> Simplifier.full_simp_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 923 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 924 | 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 | 925 |        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 | 926 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 927 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 928 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 929 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 930 | (* Quantifiers *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 931 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 932 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 933 |   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 | 934 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 935 | 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 | 936 |     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 | 937 | 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 | 938 | 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 | 939 | 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 | 940 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 941 | 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 | 942 | 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 | 943 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 944 | 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 | 945 | 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 | 946 | 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 | 947 |     addsimps @{thms HOL.simp_thms HOL.all_simps}
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 948 |     addsimprocs [@{simproc HOL.defined_All}, @{simproc HOL.defined_Ex}])
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 949 | 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 | 950 | 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 | 951 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 952 | 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 | 953 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 954 | 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 | 955 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 956 | 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 | 957 | partial_simplify_tac ctxt []) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 958 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 959 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 960 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 961 | (* Equality *) | 
| 
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 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 | 964 |   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 | 965 |   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 | 966 |   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 | 967 |   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 | 968 |   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 | 969 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 970 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 971 | 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 | 972 | (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 | 973 |       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 | 974 | | NONE => | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 975 | (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 | 976 | SOME thm => SOME thm | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 977 | | NONE => NONE)) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 978 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 979 | 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 | 980 | (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 | 981 | 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 | 982 | | NONE => | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 983 | (case t of | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 984 | f $ arg => | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 985 | (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 | 986 | 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 | 987 | | 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 | 988 | | _ => Conv.all_conv ctrm)) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 989 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 990 |   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 | 991 | let | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 992 | val rews = prems | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 993 | |> 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 | 994 | 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 | 995 |        |> 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 | 996 | 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 | 997 | 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 | 998 | val (t1, conv_term) = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 999 | (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 | 1000 | 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 | 1001 | | 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 | 1002 | | 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 | 1003 | | t1 => (t1, conv_left)) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1004 | 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 | 1005 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1006 | 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 | 1007 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1008 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1009 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1010 | 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 | 1011 | 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 | 1012 |    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 | 1013 | 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 | 1014 |    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 | 1015 | ORELSE' assume_tac ctxt)) | 
| 
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 | 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 | 1018 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1019 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1020 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1021 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1022 | (* Subproof *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1023 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1024 | 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 | 1025 | 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 | 1026 |       (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 | 1027 |            @{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 | 1028 | 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 | 1029 | 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 | 1030 | 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 | 1031 | 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 | 1032 | | subproof ctxt prems t = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1033 | 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 | 1034 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1035 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1036 | (* Simplifying Steps *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1037 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1038 | (* 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 | 1039 | 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 | 1040 | 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 | 1041 | *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1042 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1043 | 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 | 1044 | ctxt | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1045 | |> empty_simpset | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1046 | |> 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 | 1047 | |> (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 | 1048 | |> Simplifier.full_simp_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1049 | 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 | 1050 | 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 | 1051 | 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 | 1052 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1053 | 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 | 1054 | ctxt | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1055 | |> empty_simpset | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1056 | |> 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 | 1057 |     |> Raw_Simplifier.add_eqcong @{thm eq_reflection[OF imp_cong]}
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1058 |     |> (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 | 1059 | |> Simplifier.full_simp_tac | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1060 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1061 | 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 | 1062 | 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 | 1063 | 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 | 1064 | 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 | 1065 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1066 | 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 | 1067 | 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 | 1068 | 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 | 1069 | 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 | 1070 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1071 | 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 | 1072 | 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 | 1073 | 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 | 1074 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1075 |   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 | 1076 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1077 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1078 | 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 | 1079 | 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 | 1080 |    (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 | 1081 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1082 | 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 | 1083 | 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 | 1084 |    (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 | 1085 |      @{thms verit_and_simplify}))
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1086 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1087 | 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 | 1088 | 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 | 1089 |    (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 | 1090 |     @{thms verit_or_simplify})
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1091 |     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 | 1092 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1093 | 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 | 1094 | 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 | 1095 |    (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 | 1096 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1097 | 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 | 1098 | 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 | 1099 |    (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 | 1100 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1101 | 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 | 1102 | 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 | 1103 | (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 | 1104 |       @{thms verit_eq_simplify}
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1105 |       (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 | 1106 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1107 | 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 | 1108 | 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 | 1109 |     (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 | 1110 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1111 | (* 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 | 1112 | 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 | 1113 | 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 | 1114 | other direction. *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1115 | 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 | 1116 | 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 | 1117 | 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 | 1118 |         (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 | 1119 |         (@{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 | 1120 |         (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 | 1121 |         (@{thms verit_connective_def}))
 | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1122 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1123 | 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 | 1124 | 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 | 1125 | 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 | 1126 |          @{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 | 1127 |          (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 | 1128 |           @{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 | 1129 | 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 | 1130 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1131 | 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 | 1132 | 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 | 1133 |         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 | 1134 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1135 | 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 | 1136 | 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 | 1137 |    (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 | 1138 | end | 
| 
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 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1141 | (* Simplifications *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1142 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1143 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1144 | 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 | 1145 |     (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 | 1146 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1147 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1148 | 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 | 1149 | 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 | 1150 |     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 | 1151 | 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 | 1152 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1153 | 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 | 1154 | 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 | 1155 |     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 | 1156 | 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 | 1157 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1158 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1159 | 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 | 1160 | 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 | 1161 | 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 | 1162 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1163 | (* Arithmetics *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1164 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1165 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1166 | 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 | 1167 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1168 | 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 | 1169 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1170 | 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 | 1171 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1172 | 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 | 1173 | 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 | 1174 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1175 | 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 | 1176 | 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 | 1177 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1178 | 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 | 1179 |   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 | 1180 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1181 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1182 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1183 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1184 | (* Symm and Not_Symm*) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1185 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1186 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1187 | 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 | 1188 | 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 | 1189 | 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 | 1190 | | 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 | 1191 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1192 | val symm_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1193 |   @{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 | 1194 | 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 | 1195 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1196 | val not_symm_thm = | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1197 |   @{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 | 1198 | 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 | 1199 | end | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1200 | |
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1201 | (* Reordering *) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1202 | local | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1203 | 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 | 1204 | 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 | 1205 |       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 | 1206 |       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 | 1207 |       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 | 1208 |       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 | 1209 |       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 | 1210 | )) | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1211 | | 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 | 1212 | 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 | 1213 | in | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1214 | val reordering = prove_reordering | 
| 
da591621d6ae
split veriT reconstruction into Lethe and veriT part
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: diff
changeset | 1215 | end | 
| 
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 | end; |