| author | desharna | 
| Mon, 19 Dec 2022 15:54:03 +0100 | |
| changeset 76748 | b35ffbe82031 | 
| parent 69593 | 3dda49e08b9d | 
| child 79399 | 11b53e039f6f | 
| permissions | -rw-r--r-- | 
| 58061 | 1 | (* Title: HOL/Tools/SMT/z3_proof.ML | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 2 | Author: Sascha Boehme, TU Muenchen | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 3 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 4 | Z3 proofs: parsing and abstract syntax tree. | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 5 | *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 6 | |
| 58061 | 7 | signature Z3_PROOF = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 8 | sig | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 9 | (*proof rules*) | 
| 57746 | 10 | datatype z3_rule = | 
| 11 | True_Axiom | Asserted | Goal | Modus_Ponens | Reflexivity | Symmetry | Transitivity | | |
| 12 | Transitivity_Star | Monotonicity | Quant_Intro | Distributivity | And_Elim | Not_Or_Elim | | |
| 13 | Rewrite | Rewrite_Star | Pull_Quant | Pull_Quant_Star | Push_Quant | Elim_Unused_Vars | | |
| 14 | Dest_Eq_Res | Quant_Inst | Hypothesis | Lemma | Unit_Resolution | Iff_True | Iff_False | | |
| 15 | Commutativity | Def_Axiom | Intro_Def | Apply_Def | Iff_Oeq | Nnf_Pos | Nnf_Neg | Nnf_Star | | |
| 16 | Cnf_Star | Skolemize | Modus_Ponens_Oeq | Th_Lemma of string | |
| 17 | ||
| 57220 | 18 | val is_assumption: z3_rule -> bool | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 19 | val string_of_rule: z3_rule -> string | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 20 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 21 | (*proofs*) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 22 |   datatype z3_step = Z3_Step of {
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 23 | id: int, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 24 | rule: z3_rule, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 25 | prems: int list, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 26 | concl: term, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 27 | fixes: string list, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 28 | is_fix_step: bool} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 29 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 30 | (*proof parser*) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 31 | val parse: typ Symtab.table -> term Symtab.table -> string list -> | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 32 | Proof.context -> z3_step list * Proof.context | 
| 57219 
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
 blanchet parents: 
56811diff
changeset | 33 | end; | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 34 | |
| 58061 | 35 | structure Z3_Proof: Z3_PROOF = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 36 | struct | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 37 | |
| 58061 | 38 | open SMTLIB_Proof | 
| 57219 
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
 blanchet parents: 
56811diff
changeset | 39 | |
| 
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
 blanchet parents: 
56811diff
changeset | 40 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 41 | (* proof rules *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 42 | |
| 57746 | 43 | datatype z3_rule = | 
| 44 | True_Axiom | Asserted | Goal | Modus_Ponens | Reflexivity | Symmetry | Transitivity | | |
| 45 | Transitivity_Star | Monotonicity | Quant_Intro | Distributivity | And_Elim | Not_Or_Elim | | |
| 46 | Rewrite | Rewrite_Star | Pull_Quant | Pull_Quant_Star | Push_Quant | Elim_Unused_Vars | | |
| 47 | Dest_Eq_Res | Quant_Inst | Hypothesis | Lemma | Unit_Resolution | Iff_True | Iff_False | | |
| 48 | Commutativity | Def_Axiom | Intro_Def | Apply_Def | Iff_Oeq | Nnf_Pos | Nnf_Neg | Nnf_Star | | |
| 49 | Cnf_Star | Skolemize | Modus_Ponens_Oeq | Th_Lemma of string | |
| 50 | (* some proof rules include further information that is currently dropped by the parser *) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 51 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 52 | val rule_names = Symtab.make [ | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 53 |   ("true-axiom", True_Axiom),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 54 |   ("asserted", Asserted),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 55 |   ("goal", Goal),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 56 |   ("mp", Modus_Ponens),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 57 |   ("refl", Reflexivity),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 58 |   ("symm", Symmetry),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 59 |   ("trans", Transitivity),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 60 |   ("trans*", Transitivity_Star),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 61 |   ("monotonicity", Monotonicity),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 62 |   ("quant-intro", Quant_Intro),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 63 |   ("distributivity", Distributivity),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 64 |   ("and-elim", And_Elim),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 65 |   ("not-or-elim", Not_Or_Elim),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 66 |   ("rewrite", Rewrite),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 67 |   ("rewrite*", Rewrite_Star),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 68 |   ("pull-quant", Pull_Quant),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 69 |   ("pull-quant*", Pull_Quant_Star),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 70 |   ("push-quant", Push_Quant),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 71 |   ("elim-unused", Elim_Unused_Vars),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 72 |   ("der", Dest_Eq_Res),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 73 |   ("quant-inst", Quant_Inst),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 74 |   ("hypothesis", Hypothesis),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 75 |   ("lemma", Lemma),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 76 |   ("unit-resolution", Unit_Resolution),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 77 |   ("iff-true", Iff_True),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 78 |   ("iff-false", Iff_False),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 79 |   ("commutativity", Commutativity),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 80 |   ("def-axiom", Def_Axiom),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 81 |   ("intro-def", Intro_Def),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 82 |   ("apply-def", Apply_Def),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 83 |   ("iff~", Iff_Oeq),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 84 |   ("nnf-pos", Nnf_Pos),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 85 |   ("nnf-neg", Nnf_Neg),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 86 |   ("nnf*", Nnf_Star),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 87 |   ("cnf*", Cnf_Star),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 88 |   ("sk", Skolemize),
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 89 |   ("mp~", Modus_Ponens_Oeq)]
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 90 | |
| 57220 | 91 | fun is_assumption Asserted = true | 
| 92 | | is_assumption Goal = true | |
| 93 | | is_assumption Hypothesis = true | |
| 94 | | is_assumption Intro_Def = true | |
| 95 | | is_assumption Skolemize = true | |
| 96 | | is_assumption _ = false | |
| 97 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 98 | fun rule_of_string name = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 99 | (case Symtab.lookup rule_names name of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 100 | SOME rule => rule | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 101 |   | NONE => error ("unknown Z3 proof rule " ^ quote name))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 102 | |
| 57762 | 103 | fun string_of_rule (Th_Lemma kind) = "th-lemma" ^ (if kind = "" then "" else " " ^ kind) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 104 | | string_of_rule r = | 
| 57219 
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
 blanchet parents: 
56811diff
changeset | 105 | let fun eq_rule (s, r') = if r = r' then SOME s else NONE | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 106 | in the (Symtab.get_first eq_rule rule_names) end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 107 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 108 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 109 | (* proofs *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 110 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 111 | datatype z3_node = Z3_Node of {
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 112 | id: int, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 113 | rule: z3_rule, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 114 | prems: z3_node list, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 115 | concl: term, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 116 | bounds: string list} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 117 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 118 | fun mk_node id rule prems concl bounds = | 
| 57746 | 119 |   Z3_Node {id = id, rule = rule, prems = prems, concl = concl, bounds = bounds}
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 120 | |
| 66427 | 121 | fun string_of_node ctxt = | 
| 122 | let | |
| 123 |     fun str depth (Z3_Node {id, rule, prems, concl, bounds}) =
 | |
| 124 | replicate_string depth " " ^ | |
| 125 |       enclose "{" "}" (commas
 | |
| 126 | [string_of_int id, | |
| 127 | string_of_rule rule, | |
| 128 | enclose "[" "]" (space_implode " " bounds), | |
| 129 | Syntax.string_of_term ctxt concl] ^ | |
| 130 | cat_lines (map (prefix "\n" o str (depth + 1)) prems)) | |
| 131 | in str 0 end | |
| 132 | ||
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 133 | datatype z3_step = Z3_Step of {
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 134 | id: int, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 135 | rule: z3_rule, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 136 | prems: int list, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 137 | concl: term, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 138 | fixes: string list, | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 139 | is_fix_step: bool} | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 140 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 141 | fun mk_step id rule prems concl fixes is_fix_step = | 
| 57746 | 142 |   Z3_Step {id = id, rule = rule, prems = prems, concl = concl, fixes = fixes,
 | 
| 143 | is_fix_step = is_fix_step} | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 144 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 145 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 146 | (* proof parser *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 147 | |
| 58061 | 148 | fun rule_of (SMTLIB.Sym name) = rule_of_string name | 
| 149 | | rule_of (SMTLIB.S (SMTLIB.Sym "_" :: SMTLIB.Sym name :: args)) = | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 150 | (case (name, args) of | 
| 58061 | 151 |         ("th-lemma", SMTLIB.Sym kind :: _) => Th_Lemma kind
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 152 | | _ => rule_of_string name) | 
| 58061 | 153 |   | rule_of r = raise SMTLIB_PARSE ("bad Z3 proof rule format", r)
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 154 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 155 | fun node_of p cx = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 156 | (case p of | 
| 58061 | 157 | SMTLIB.Sym name => | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 158 | (case lookup_binding cx name of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 159 | Proof node => (node, cx) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 160 | | Tree p' => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 161 | cx | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 162 | |> node_of p' | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 163 | |-> (fn node => pair node o update_binding (name, Proof node)) | 
| 58061 | 164 |       | _ => raise SMTLIB_PARSE ("bad Z3 proof format", p))
 | 
| 165 | | SMTLIB.S [SMTLIB.Sym "let", SMTLIB.S bindings, p] => | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 166 | with_bindings (map dest_binding bindings) (node_of p) cx | 
| 58061 | 167 | | SMTLIB.S (name :: parts) => | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 168 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 169 | val (ps, p) = split_last parts | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 170 | val r = rule_of name | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 171 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 172 | cx | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 173 | |> fold_map node_of ps | 
| 57747 | 174 | ||>> `(with_fresh_names (term_of p)) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 175 | ||>> next_id | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 176 | |>> (fn ((prems, (t, ns)), id) => mk_node id r prems t ns) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 177 | end | 
| 58061 | 178 |   | _ => raise SMTLIB_PARSE ("bad Z3 proof format", p))
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 179 | |
| 58061 | 180 | fun dest_name (SMTLIB.Sym name) = name | 
| 181 |   | dest_name t = raise SMTLIB_PARSE ("bad name", t)
 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 182 | |
| 58061 | 183 | fun dest_seq (SMTLIB.S ts) = ts | 
| 184 |   | dest_seq t = raise SMTLIB_PARSE ("bad Z3 proof format", t)
 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 185 | |
| 58061 | 186 | fun parse' (SMTLIB.S (SMTLIB.Sym "set-logic" :: _) :: ts) cx = parse' ts cx | 
| 187 | | parse' (SMTLIB.S [SMTLIB.Sym "declare-fun", n, tys, ty] :: ts) cx = | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 188 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 189 | val name = dest_name n | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 190 | val Ts = map (type_of cx) (dest_seq tys) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 191 | val T = type_of cx ty | 
| 57219 
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
 blanchet parents: 
56811diff
changeset | 192 | in parse' ts (declare_fun name (Ts ---> T) cx) end | 
| 58061 | 193 | | parse' (SMTLIB.S [SMTLIB.Sym "proof", p] :: _) cx = node_of p cx | 
| 194 |   | parse' ts _ = raise SMTLIB_PARSE ("bad Z3 proof declarations", SMTLIB.S ts)
 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 195 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 196 | fun parse_proof typs funs lines ctxt = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 197 | let | 
| 58061 | 198 | val ts = dest_seq (SMTLIB.parse lines) | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 199 | val (node, cx) = parse' ts (empty_context ctxt typs funs) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 200 | in (node, ctxt_of cx) end | 
| 58061 | 201 |   handle SMTLIB.PARSE (l, msg) => error ("parsing error at line " ^ string_of_int l ^ ": " ^ msg)
 | 
| 202 | | SMTLIB_PARSE (msg, t) => error (msg ^ ": " ^ SMTLIB.str_of t) | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 203 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 204 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 205 | (* handling of bound variables *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 206 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 207 | fun subst_of tyenv = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 208 | let fun add (ix, (S, T)) = cons (TVar (ix, S), T) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 209 | in Vartab.fold add tyenv [] end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 210 | |
| 57219 
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
 blanchet parents: 
56811diff
changeset | 211 | fun substTs_same subst = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 212 | let val applyT = Same.function (AList.lookup (op =) subst) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 213 | in Term_Subst.map_atypsT_same applyT end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 214 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 215 | fun subst_types ctxt env bounds t = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 216 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 217 | val match = Sign.typ_match (Proof_Context.theory_of ctxt) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 218 | |
| 68057 
7811d8828775
prefer explicit error message to unspecific Options exception: Z3 proof traces may lack information necessary for replay when dealing with quantified formulas
 boehmes parents: 
66427diff
changeset | 219 | fun objT_of bound = | 
| 
7811d8828775
prefer explicit error message to unspecific Options exception: Z3 proof traces may lack information necessary for replay when dealing with quantified formulas
 boehmes parents: 
66427diff
changeset | 220 | (case Symtab.lookup env bound of | 
| 
7811d8828775
prefer explicit error message to unspecific Options exception: Z3 proof traces may lack information necessary for replay when dealing with quantified formulas
 boehmes parents: 
66427diff
changeset | 221 | SOME objT => objT | 
| 
7811d8828775
prefer explicit error message to unspecific Options exception: Z3 proof traces may lack information necessary for replay when dealing with quantified formulas
 boehmes parents: 
66427diff
changeset | 222 |       | NONE => raise Fail ("Replaying the proof trace produced by Z3 failed: " ^
 | 
| 
7811d8828775
prefer explicit error message to unspecific Options exception: Z3 proof traces may lack information necessary for replay when dealing with quantified formulas
 boehmes parents: 
66427diff
changeset | 223 | "the bound " ^ quote bound ^ " is undeclared; this indicates a bug in Z3")) | 
| 
7811d8828775
prefer explicit error message to unspecific Options exception: Z3 proof traces may lack information necessary for replay when dealing with quantified formulas
 boehmes parents: 
66427diff
changeset | 224 | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 225 | val t' = singleton (Variable.polymorphic ctxt) t | 
| 69593 | 226 | val patTs = map snd (Term.strip_qnt_vars \<^const_name>\<open>Pure.all\<close> t') | 
| 68057 
7811d8828775
prefer explicit error message to unspecific Options exception: Z3 proof traces may lack information necessary for replay when dealing with quantified formulas
 boehmes parents: 
66427diff
changeset | 227 | val objTs = map objT_of bounds | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 228 | val subst = subst_of (fold match (patTs ~~ objTs) Vartab.empty) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 229 | in Same.commit (Term_Subst.map_types_same (substTs_same subst)) t' end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 230 | |
| 69593 | 231 | fun eq_quant (\<^const_name>\<open>HOL.All\<close>, _) (\<^const_name>\<open>HOL.All\<close>, _) = true | 
| 232 | | eq_quant (\<^const_name>\<open>HOL.Ex\<close>, _) (\<^const_name>\<open>HOL.Ex\<close>, _) = true | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 233 | | eq_quant _ _ = false | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 234 | |
| 69593 | 235 | fun opp_quant (\<^const_name>\<open>HOL.All\<close>, _) (\<^const_name>\<open>HOL.Ex\<close>, _) = true | 
| 236 | | opp_quant (\<^const_name>\<open>HOL.Ex\<close>, _) (\<^const_name>\<open>HOL.All\<close>, _) = true | |
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 237 | | opp_quant _ _ = false | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 238 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 239 | fun with_quant pred i (Const q1 $ Abs (_, T1, t1), Const q2 $ Abs (_, T2, t2)) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 240 | if pred q1 q2 andalso T1 = T2 then | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 241 |         let val t = Var (("", i), T1)
 | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58061diff
changeset | 242 | in SOME (apply2 Term.subst_bound ((t, t1), (t, t2))) end | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 243 | else NONE | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 244 | | with_quant _ _ _ = NONE | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 245 | |
| 69593 | 246 | fun dest_quant_pair i (\<^term>\<open>HOL.Not\<close> $ t1, t2) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 247 | Option.map (apfst HOLogic.mk_not) (with_quant opp_quant i (t1, t2)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 248 | | dest_quant_pair i (t1, t2) = with_quant eq_quant i (t1, t2) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 249 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 250 | fun dest_quant i t = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 251 | (case dest_quant_pair i (HOLogic.dest_eq (HOLogic.dest_Trueprop t)) of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 252 | SOME (t1, t2) => HOLogic.mk_Trueprop (HOLogic.mk_eq (t1, t2)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 253 |   | NONE => raise TERM ("lift_quant", [t]))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 254 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 255 | fun match_types ctxt pat obj = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 256 | (Vartab.empty, Vartab.empty) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 257 | |> Pattern.first_order_match (Proof_Context.theory_of ctxt) (pat, obj) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 258 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 259 | fun strip_match ctxt pat i obj = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 260 | (case try (match_types ctxt pat) obj of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 261 | SOME (tyenv, _) => subst_of tyenv | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 262 | | NONE => strip_match ctxt pat (i + 1) (dest_quant i obj)) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 263 | |
| 69593 | 264 | fun dest_all i (Const (\<^const_name>\<open>Pure.all\<close>, _) $ (a as Abs (_, T, _))) = | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 265 |       dest_all (i + 1) (Term.betapply (a, Var (("", i), T)))
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 266 | | dest_all i t = (i, t) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 267 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 268 | fun dest_alls t = dest_all (Term.maxidx_of_term t + 1) t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 269 | |
| 57746 | 270 | fun match_rule ctxt env (Z3_Node {bounds = bs', concl = t', ...}) bs t =
 | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 271 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 272 | val t'' = singleton (Variable.polymorphic ctxt) t' | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 273 | val (i, obj) = dest_alls (subst_types ctxt env bs t) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 274 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 275 | (case try (strip_match ctxt (snd (dest_alls t'')) i) obj of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 276 | NONE => NONE | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 277 | | SOME subst => | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 278 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 279 | val applyT = Same.commit (substTs_same subst) | 
| 69593 | 280 | val patTs = map snd (Term.strip_qnt_vars \<^const_name>\<open>Pure.all\<close> t'') | 
| 56078 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 281 | in SOME (Symtab.make (bs' ~~ map applyT patTs)) end) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 282 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 283 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 284 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 285 | (* linearizing proofs and resolving types of bound variables *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 286 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 287 | fun has_step (tab, _) = Inttab.defined tab | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 288 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 289 | fun add_step id rule bounds concl is_fix_step ids (tab, sts) = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 290 | let val step = mk_step id rule ids concl bounds is_fix_step | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 291 | in (id, (Inttab.update (id, ()) tab, step :: sts)) end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 292 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 293 | fun is_fix_rule rule prems = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 294 | member (op =) [Quant_Intro, Nnf_Pos, Nnf_Neg] rule andalso length prems = 1 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 295 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 296 | fun lin_proof ctxt env (Z3_Node {id, rule, prems, concl, bounds}) steps =
 | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 297 | if has_step steps id then (id, steps) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 298 | else | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 299 | let | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 300 | val t = subst_types ctxt env bounds concl | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 301 | val add = add_step id rule bounds t | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 302 | fun rec_apply e b = fold_map (lin_proof ctxt e) prems #-> add b | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 303 | in | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 304 | if is_fix_rule rule prems then | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 305 | (case match_rule ctxt env (hd prems) bounds t of | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 306 | NONE => rec_apply env false steps | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 307 | | SOME env' => rec_apply env' true steps) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 308 | else rec_apply env false steps | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 309 | end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 310 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 311 | fun linearize ctxt node = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 312 | rev (snd (snd (lin_proof ctxt Symtab.empty node (Inttab.empty, [])))) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 313 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 314 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 315 | (* overall proof parser *) | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 316 | |
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 317 | fun parse typs funs lines ctxt = | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 318 | let val (node, ctxt') = parse_proof typs funs lines ctxt | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 319 | in (linearize ctxt' node, ctxt') end | 
| 
624faeda77b5
moved 'SMT2' (SMT-LIB-2-based SMT module) into Isabelle
 blanchet parents: diff
changeset | 320 | |
| 57219 
34018603e0d0
factor out SMT-LIB 2 type/term parsing from Z3-specific code
 blanchet parents: 
56811diff
changeset | 321 | end; |