| author | wenzelm | 
| Wed, 17 Feb 2016 23:29:35 +0100 | |
| changeset 62357 | ab76bd43c14a | 
| parent 61268 | abe08fb15a12 | 
| child 67091 | 1393c2340eec | 
| permissions | -rw-r--r-- | 
| 39941 | 1 | (* Title: HOL/Tools/Meson/meson_clausify.ML | 
| 38027 | 2 | Author: Jia Meng, Cambridge University Computer Laboratory and NICTA | 
| 36393 
be73a2b2443b
support readable names even when Isar proof reconstruction is enabled -- useful for debugging
 blanchet parents: 
36228diff
changeset | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 15347 | 4 | |
| 39941 | 5 | Transformation of HOL theorems into CNF forms. | 
| 15347 | 6 | *) | 
| 7 | ||
| 39890 | 8 | signature MESON_CLAUSIFY = | 
| 21505 | 9 | sig | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 10 | val new_skolem_var_prefix : string | 
| 42098 
f978caf60bbe
more robust handling of variables in new Skolemizer
 blanchet parents: 
42072diff
changeset | 11 | val new_nonskolem_var_prefix : string | 
| 42099 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 12 | val is_zapped_var_name : string -> bool | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 13 | val is_quasi_lambda_free : term -> bool | 
| 60328 | 14 | val introduce_combinators_in_cterm : Proof.context -> cterm -> thm | 
| 55236 | 15 | val introduce_combinators_in_theorem : Proof.context -> thm -> thm | 
| 39932 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 16 | val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50705diff
changeset | 17 | val ss_only : thm list -> Proof.context -> Proof.context | 
| 39897 | 18 | val cnf_axiom : | 
| 45508 | 19 | Proof.context -> bool -> bool -> int -> thm | 
| 20 | -> (thm * term) option * thm list | |
| 21505 | 21 | end; | 
| 19196 
62ee8c10d796
Added functions to retrieve local and global atpset rules.
 mengj parents: 
19175diff
changeset | 22 | |
| 39890 | 23 | structure Meson_Clausify : MESON_CLAUSIFY = | 
| 15997 | 24 | struct | 
| 15347 | 25 | |
| 39950 | 26 | open Meson | 
| 27 | ||
| 42072 
1492ee6b8085
avoid illegal variable names, which then yield "Not an identifier" errors in "rename_tac"
 blanchet parents: 
41225diff
changeset | 28 | (* the extra "Meson" helps prevent clashes (FIXME) *) | 
| 
1492ee6b8085
avoid illegal variable names, which then yield "Not an identifier" errors in "rename_tac"
 blanchet parents: 
41225diff
changeset | 29 | val new_skolem_var_prefix = "MesonSK" | 
| 
1492ee6b8085
avoid illegal variable names, which then yield "Not an identifier" errors in "rename_tac"
 blanchet parents: 
41225diff
changeset | 30 | val new_nonskolem_var_prefix = "MesonV" | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 31 | |
| 42099 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 32 | fun is_zapped_var_name s = | 
| 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 33 | exists (fn prefix => String.isPrefix prefix s) | 
| 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 34 | [new_skolem_var_prefix, new_nonskolem_var_prefix] | 
| 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 35 | |
| 15997 | 36 | (**** Transformation of Elimination Rules into First-Order Formulas****) | 
| 15347 | 37 | |
| 59632 | 38 | val cfalse = Thm.cterm_of @{theory_context HOL} @{term False};
 | 
| 39 | val ctp_false = Thm.cterm_of @{theory_context HOL} (HOLogic.mk_Trueprop @{term False});
 | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 40 | |
| 38001 
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
 blanchet parents: 
38000diff
changeset | 41 | (* Converts an elim-rule into an equivalent theorem that does not have the | 
| 
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
 blanchet parents: 
38000diff
changeset | 42 | predicate variable. Leaves other theorems unchanged. We simply instantiate | 
| 42944 
9e620869a576
improved Waldmeister support -- even run it by default on unit equational goals
 blanchet parents: 
42747diff
changeset | 43 | the conclusion variable to False. (Cf. "transform_elim_prop" in | 
| 38652 
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
 blanchet parents: 
38632diff
changeset | 44 | "Sledgehammer_Util".) *) | 
| 38001 
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
 blanchet parents: 
38000diff
changeset | 45 | fun transform_elim_theorem th = | 
| 59582 | 46 | (case Thm.concl_of th of (*conclusion variable*) | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60328diff
changeset | 47 |     @{const Trueprop} $ (Var (v as (_, @{typ bool}))) =>
 | 
| 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60328diff
changeset | 48 | Thm.instantiate ([], [(v, cfalse)]) th | 
| 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60328diff
changeset | 49 |   | Var (v as (_, @{typ prop})) =>
 | 
| 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60328diff
changeset | 50 | Thm.instantiate ([], [(v, ctp_false)]) th | 
| 59582 | 51 | | _ => th) | 
| 15997 | 52 | |
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 53 | |
| 16009 | 54 | (**** SKOLEMIZATION BY INFERENCE (lcp) ****) | 
| 55 | ||
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39721diff
changeset | 56 | fun mk_old_skolem_term_wrapper t = | 
| 37436 | 57 | let val T = fastype_of t in | 
| 39962 
d42ddd7407ca
qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
 blanchet parents: 
39950diff
changeset | 58 |     Const (@{const_name Meson.skolem}, T --> T) $ t
 | 
| 37436 | 59 | end | 
| 37410 
2bf7e6136047
adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
 blanchet parents: 
37403diff
changeset | 60 | |
| 39931 | 61 | fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') | 
| 62 | | beta_eta_in_abs_body t = Envir.beta_eta_contract t | |
| 37512 
ff72d3ddc898
this looks like the most appropriate place to do the beta-eta-contraction
 blanchet parents: 
37511diff
changeset | 63 | |
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 64 | (*Traverse a theorem, accumulating Skolem function definitions.*) | 
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39721diff
changeset | 65 | fun old_skolem_defs th = | 
| 37399 
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
 blanchet parents: 
37349diff
changeset | 66 | let | 
| 39376 | 67 |     fun dec_sko (Const (@{const_name Ex}, _) $ (body as Abs (_, T, p))) rhss =
 | 
| 37399 
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
 blanchet parents: 
37349diff
changeset | 68 | (*Existential: declare a Skolem function, then insert into body and continue*) | 
| 
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
 blanchet parents: 
37349diff
changeset | 69 | let | 
| 44121 | 70 | val args = Misc_Legacy.term_frees body | 
| 37500 
7587b6e63454
thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
 blanchet parents: 
37498diff
changeset | 71 | (* Forms a lambda-abstraction over the formal parameters *) | 
| 
7587b6e63454
thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
 blanchet parents: 
37498diff
changeset | 72 | val rhs = | 
| 44241 | 73 | fold_rev (absfree o dest_Free) args | 
| 74 | (HOLogic.choice_const T $ beta_eta_in_abs_body body) | |
| 39886 
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
 blanchet parents: 
39721diff
changeset | 75 | |> mk_old_skolem_term_wrapper | 
| 37518 
efb0923cc098
use Skolem inlining also in the nonpolymorphic case, as a step toward simplifying the source code
 blanchet parents: 
37512diff
changeset | 76 | val comb = list_comb (rhs, args) | 
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 77 | in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end | 
| 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 78 |       | dec_sko (Const (@{const_name All},_) $ Abs (a, T, p)) rhss =
 | 
| 37399 
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
 blanchet parents: 
37349diff
changeset | 79 | (*Universal quant: insert a free variable into body and continue*) | 
| 44121 | 80 | let val fname = singleton (Name.variant_list (Misc_Legacy.add_term_names (p, []))) a | 
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 81 | in dec_sko (subst_bound (Free(fname,T), p)) rhss end | 
| 39906 | 82 |       | dec_sko (@{const conj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
 | 
| 83 |       | dec_sko (@{const disj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q
 | |
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 84 |       | dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss
 | 
| 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 85 | | dec_sko _ rhss = rhss | 
| 59582 | 86 | in dec_sko (Thm.prop_of th) [] end; | 
| 20419 | 87 | |
| 88 | ||
| 24827 | 89 | (**** REPLACING ABSTRACTIONS BY COMBINATORS ****) | 
| 20419 | 90 | |
| 39962 
d42ddd7407ca
qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
 blanchet parents: 
39950diff
changeset | 91 | fun is_quasi_lambda_free (Const (@{const_name Meson.skolem}, _) $ _) = true
 | 
| 37416 
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
 blanchet parents: 
37410diff
changeset | 92 | | is_quasi_lambda_free (t1 $ t2) = | 
| 
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
 blanchet parents: 
37410diff
changeset | 93 | is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2 | 
| 
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
 blanchet parents: 
37410diff
changeset | 94 | | is_quasi_lambda_free (Abs _) = false | 
| 
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
 blanchet parents: 
37410diff
changeset | 95 | | is_quasi_lambda_free _ = true | 
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 96 | |
| 38282 
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
 blanchet parents: 
38280diff
changeset | 97 | (* FIXME: Requires more use of cterm constructors. *) | 
| 60328 | 98 | fun abstract ctxt ct = | 
| 28544 
26743a1591f5
improved performance of skolem cache, due to parallel map;
 wenzelm parents: 
28262diff
changeset | 99 | let | 
| 59582 | 100 | val Abs(x,_,body) = Thm.term_of ct | 
| 59586 | 101 |       val Type (@{type_name fun}, [xT,bodyT]) = Thm.typ_of_cterm ct
 | 
| 60328 | 102 | val cxT = Thm.ctyp_of ctxt xT | 
| 103 | val cbodyT = Thm.ctyp_of ctxt bodyT | |
| 38005 
b6555e9c5de4
prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
 blanchet parents: 
38001diff
changeset | 104 | fun makeK () = | 
| 60801 | 105 |         Thm.instantiate' [SOME cxT, SOME cbodyT] [SOME (Thm.cterm_of ctxt body)] @{thm abs_K}
 | 
| 24827 | 106 | in | 
| 107 | case body of | |
| 108 | Const _ => makeK() | |
| 109 | | Free _ => makeK() | |
| 110 | | Var _ => makeK() (*though Var isn't expected*) | |
| 60801 | 111 |         | Bound 0 => Thm.instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*)
 | 
| 24827 | 112 | | rator$rand => | 
| 42083 
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
 wenzelm parents: 
42072diff
changeset | 113 | if Term.is_dependent rator then (*C or S*) | 
| 
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
 wenzelm parents: 
42072diff
changeset | 114 | if Term.is_dependent rand then (*S*) | 
| 60328 | 115 | let val crator = Thm.cterm_of ctxt (Abs (x, xT, rator)) | 
| 116 | val crand = Thm.cterm_of ctxt (Abs (x, xT, rand)) | |
| 60781 | 117 | val abs_S' = | 
| 118 |                       infer_instantiate ctxt [(("f", 0), crator), (("g", 0), crand)] @{thm abs_S}
 | |
| 59582 | 119 | val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_S') | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 120 | in | 
| 60328 | 121 | Thm.transitive abs_S' (Conv.binop_conv (abstract ctxt) rhs) | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 122 | end | 
| 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 123 | else (*C*) | 
| 60328 | 124 | let val crator = Thm.cterm_of ctxt (Abs (x, xT, rator)) | 
| 59582 | 125 | val abs_C' = | 
| 60781 | 126 |                       infer_instantiate ctxt [(("f", 0), crator), (("b", 0), Thm.cterm_of ctxt rand)]
 | 
| 127 |                         @{thm abs_C}
 | |
| 59582 | 128 | val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_C') | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 129 | in | 
| 60328 | 130 | Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv (abstract ctxt)) rhs) | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 131 | end | 
| 42083 
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
 wenzelm parents: 
42072diff
changeset | 132 | else if Term.is_dependent rand then (*B or eta*) | 
| 36945 | 133 | if rand = Bound 0 then Thm.eta_conversion ct | 
| 27179 
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
 wenzelm parents: 
27174diff
changeset | 134 | else (*B*) | 
| 60328 | 135 | let val crand = Thm.cterm_of ctxt (Abs (x, xT, rand)) | 
| 136 | val crator = Thm.cterm_of ctxt rator | |
| 60781 | 137 | val abs_B' = | 
| 138 |                       infer_instantiate ctxt [(("a", 0), crator), (("g", 0), crand)] @{thm abs_B}
 | |
| 59582 | 139 | val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_B') | 
| 60328 | 140 | in Thm.transitive abs_B' (Conv.arg_conv (abstract ctxt) rhs) end | 
| 141 | else makeK () | |
| 37349 | 142 | | _ => raise Fail "abstract: Bad term" | 
| 24827 | 143 | end; | 
| 20863 
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
 paulson parents: 
20789diff
changeset | 144 | |
| 37349 | 145 | (* Traverse a theorem, remplacing lambda-abstractions with combinators. *) | 
| 60328 | 146 | fun introduce_combinators_in_cterm ctxt ct = | 
| 59582 | 147 | if is_quasi_lambda_free (Thm.term_of ct) then | 
| 37349 | 148 | Thm.reflexive ct | 
| 59582 | 149 | else case Thm.term_of ct of | 
| 37349 | 150 | Abs _ => | 
| 151 | let | |
| 152 | val (cv, cta) = Thm.dest_abs NONE ct | |
| 59582 | 153 | val (v, _) = dest_Free (Thm.term_of cv) | 
| 60328 | 154 | val u_th = introduce_combinators_in_cterm ctxt cta | 
| 37349 | 155 | val cu = Thm.rhs_of u_th | 
| 60328 | 156 | val comb_eq = abstract ctxt (Thm.lambda cv cu) | 
| 37349 | 157 | in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end | 
| 158 | | _ $ _ => | |
| 159 | let val (ct1, ct2) = Thm.dest_comb ct in | |
| 60328 | 160 | Thm.combination (introduce_combinators_in_cterm ctxt ct1) | 
| 161 | (introduce_combinators_in_cterm ctxt ct2) | |
| 37349 | 162 | end | 
| 163 | ||
| 55236 | 164 | fun introduce_combinators_in_theorem ctxt th = | 
| 59582 | 165 | if is_quasi_lambda_free (Thm.prop_of th) then | 
| 37349 | 166 | th | 
| 24827 | 167 | else | 
| 37349 | 168 | let | 
| 169 | val th = Drule.eta_contraction_rule th | |
| 60328 | 170 | val eqth = introduce_combinators_in_cterm ctxt (Thm.cprop_of th) | 
| 37349 | 171 | in Thm.equal_elim eqth th end | 
| 172 | handle THM (msg, _, _) => | |
| 61268 | 173 |            (warning ("Error in the combinator translation of " ^ Thm.string_of_thm ctxt th ^
 | 
| 55523 
9429e7b5b827
removed final periods in messages for proof methods
 blanchet parents: 
55239diff
changeset | 174 | "\nException message: " ^ msg); | 
| 45511 
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
 blanchet parents: 
45508diff
changeset | 175 |             (* A type variable of sort "{}" will make "abstraction" fail. *)
 | 
| 37349 | 176 | TrueI) | 
| 16009 | 177 | |
| 178 | (*cterms are used throughout for efficiency*) | |
| 59632 | 179 | val cTrueprop = Thm.cterm_of @{theory_context HOL} HOLogic.Trueprop;
 | 
| 16009 | 180 | |
| 181 | (*Given an abstraction over n variables, replace the bound variables by free | |
| 182 | ones. Return the body, along with the list of free variables.*) | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 183 | fun c_variant_abs_multi (ct0, vars) = | 
| 16009 | 184 | let val (cv,ct) = Thm.dest_abs NONE ct0 | 
| 185 | in c_variant_abs_multi (ct, cv::vars) end | |
| 186 | handle CTERM _ => (ct0, rev vars); | |
| 187 | ||
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 188 | (* Given the definition of a Skolem function, return a theorem to replace | 
| 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 189 | an existential formula by a use of that function. | 
| 18141 
89e2e8bed08f
Skolemization by inference, but not quite finished
 paulson parents: 
18009diff
changeset | 190 | Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *) | 
| 57466 | 191 | fun old_skolem_theorem_of_def ctxt rhs0 = | 
| 37399 
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
 blanchet parents: 
37349diff
changeset | 192 | let | 
| 59632 | 193 | val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> Thm.cterm_of ctxt | 
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 194 | val rhs' = rhs |> Thm.dest_comb |> snd | 
| 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 195 | val (ch, frees) = c_variant_abs_multi (rhs', []) | 
| 59582 | 196 | val (hilbert, cabs) = ch |> Thm.dest_comb |>> Thm.term_of | 
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 197 | val T = | 
| 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 198 | case hilbert of | 
| 39941 | 199 |         Const (_, Type (@{type_name fun}, [_, T])) => T
 | 
| 57466 | 200 |       | _ => raise TERM ("old_skolem_theorem_of_def: expected \"Eps\"", [hilbert])
 | 
| 59632 | 201 | val cex = Thm.cterm_of ctxt (HOLogic.exists_const T) | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
46071diff
changeset | 202 | val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs) | 
| 37629 | 203 | val conc = | 
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 204 | Drule.list_comb (rhs, frees) | 
| 46497 
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
 wenzelm parents: 
46071diff
changeset | 205 | |> Drule.beta_conv cabs |> Thm.apply cTrueprop | 
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 206 | fun tacf [prem] = | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52031diff
changeset | 207 |       rewrite_goals_tac ctxt @{thms skolem_def [abs_def]}
 | 
| 59632 | 208 | THEN resolve_tac ctxt | 
| 209 |         [(prem |> rewrite_rule ctxt @{thms skolem_def [abs_def]})
 | |
| 210 | RS Global_Theory.get_thm (Proof_Context.theory_of ctxt) "Hilbert_Choice.someI_ex"] 1 | |
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 211 | in | 
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 212 | Goal.prove_internal ctxt [ex_tm] conc tacf | 
| 37629 | 213 | |> forall_intr_list frees | 
| 214 | |> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) | |
| 215 | |> Thm.varifyT_global | |
| 37617 
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
 blanchet parents: 
37616diff
changeset | 216 | end | 
| 24742 
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
 paulson parents: 
24669diff
changeset | 217 | |
| 42335 | 218 | fun to_definitional_cnf_with_quantifiers ctxt th = | 
| 39036 
dff91b90d74c
use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
 blanchet parents: 
38864diff
changeset | 219 | let | 
| 59582 | 220 | val eqth = CNF.make_cnfx_thm ctxt (HOLogic.dest_Trueprop (Thm.prop_of th)) | 
| 39036 
dff91b90d74c
use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
 blanchet parents: 
38864diff
changeset | 221 |     val eqth = eqth RS @{thm eq_reflection}
 | 
| 
dff91b90d74c
use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
 blanchet parents: 
38864diff
changeset | 222 |     val eqth = eqth RS @{thm TruepropI}
 | 
| 
dff91b90d74c
use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
 blanchet parents: 
38864diff
changeset | 223 | in Thm.equal_elim eqth th end | 
| 
dff91b90d74c
use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
 blanchet parents: 
38864diff
changeset | 224 | |
| 39932 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 225 | fun zapped_var_name ((ax_no, cluster_no), skolem) index_no s = | 
| 39896 
13b3a2ba9ea7
encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
 blanchet parents: 
39894diff
changeset | 226 | (if skolem then new_skolem_var_prefix else new_nonskolem_var_prefix) ^ | 
| 39932 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 227 | "_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^ | 
| 56811 | 228 | string_of_int index_no ^ "_" ^ Name.desymbolize (SOME false) s | 
| 39896 
13b3a2ba9ea7
encode axiom number and cluster number in all zapped quantifiers to help discharging new skolemizer assumptions
 blanchet parents: 
39894diff
changeset | 229 | |
| 39899 
608b108ec979
compute quantifier dependency graph in new skolemizer
 blanchet parents: 
39897diff
changeset | 230 | fun cluster_of_zapped_var_name s = | 
| 39932 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 231 | let val get_int = the o Int.fromString o nth (space_explode "_" s) in | 
| 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 232 | ((get_int 1, (get_int 2, get_int 3)), | 
| 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 233 | String.isPrefix new_skolem_var_prefix s) | 
| 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 234 | end | 
| 39897 | 235 | |
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 236 | fun rename_bound_vars_to_be_zapped ax_no = | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 237 | let | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 238 | fun aux (cluster as (cluster_no, cluster_skolem)) index_no pos t = | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 239 | case t of | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 240 | (t1 as Const (s, _)) $ Abs (s', T, t') => | 
| 56245 | 241 |         if s = @{const_name Pure.all} orelse s = @{const_name All} orelse
 | 
| 39906 | 242 |            s = @{const_name Ex} then
 | 
| 39932 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 243 | let | 
| 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 244 |             val skolem = (pos = (s = @{const_name Ex}))
 | 
| 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 245 | val (cluster, index_no) = | 
| 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 246 | if skolem = cluster_skolem then (cluster, index_no) | 
| 
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
 blanchet parents: 
39931diff
changeset | 247 | else ((cluster_no ||> cluster_skolem ? Integer.add 1, skolem), 0) | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 248 | val s' = zapped_var_name cluster index_no s' | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 249 | in t1 $ Abs (s', T, aux cluster (index_no + 1) pos t') end | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 250 | else | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 251 | t | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 252 | | (t1 as Const (s, _)) $ t2 $ t3 => | 
| 56245 | 253 |         if s = @{const_name Pure.imp} orelse s = @{const_name HOL.implies} then
 | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 254 | t1 $ aux cluster index_no (not pos) t2 $ aux cluster index_no pos t3 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 255 |         else if s = @{const_name HOL.conj} orelse
 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 256 |                 s = @{const_name HOL.disj} then
 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 257 | t1 $ aux cluster index_no pos t2 $ aux cluster index_no pos t3 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 258 | else | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 259 | t | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 260 | | (t1 as Const (s, _)) $ t2 => | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 261 |         if s = @{const_name Trueprop} then
 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 262 | t1 $ aux cluster index_no pos t2 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 263 |         else if s = @{const_name Not} then
 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 264 | t1 $ aux cluster index_no (not pos) t2 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 265 | else | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 266 | t | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 267 | | _ => t | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 268 | in aux ((ax_no, 0), true) 0 true end | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 269 | |
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 270 | fun zap pos ct = | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 271 | ct | 
| 59582 | 272 | |> (case Thm.term_of ct of | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 273 | Const (s, _) $ Abs (s', _, _) => | 
| 56245 | 274 |         if s = @{const_name Pure.all} orelse s = @{const_name All} orelse
 | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 275 |            s = @{const_name Ex} then
 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 276 | Thm.dest_comb #> snd #> Thm.dest_abs (SOME s') #> snd #> zap pos | 
| 39906 | 277 | else | 
| 278 | Conv.all_conv | |
| 279 | | Const (s, _) $ _ $ _ => | |
| 56245 | 280 |         if s = @{const_name Pure.imp} orelse s = @{const_name implies} then
 | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 281 | Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos) | 
| 39906 | 282 |         else if s = @{const_name conj} orelse s = @{const_name disj} then
 | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 283 | Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos) | 
| 39906 | 284 | else | 
| 285 | Conv.all_conv | |
| 286 | | Const (s, _) $ _ => | |
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 287 |         if s = @{const_name Trueprop} then Conv.arg_conv (zap pos)
 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 288 |         else if s = @{const_name Not} then Conv.arg_conv (zap (not pos))
 | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 289 | else Conv.all_conv | 
| 39906 | 290 | | _ => Conv.all_conv) | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 291 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50705diff
changeset | 292 | fun ss_only ths ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps ths | 
| 39901 
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
 blanchet parents: 
39900diff
changeset | 293 | |
| 40261 
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
 blanchet parents: 
40260diff
changeset | 294 | val cheat_choice = | 
| 39901 
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
 blanchet parents: 
39900diff
changeset | 295 |   @{prop "ALL x. EX y. Q x y ==> EX f. ALL x. Q x (f x)"}
 | 
| 
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
 blanchet parents: 
39900diff
changeset | 296 | |> Logic.varify_global | 
| 
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
 blanchet parents: 
39900diff
changeset | 297 |   |> Skip_Proof.make_thm @{theory}
 | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 298 | |
| 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 299 | (* Converts an Isabelle theorem into NNF. *) | 
| 50705 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
47954diff
changeset | 300 | fun nnf_axiom choice_ths new_skolem ax_no th ctxt = | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 301 | let | 
| 42361 | 302 | val thy = Proof_Context.theory_of ctxt | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 303 | val th = | 
| 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 304 | th |> transform_elim_theorem | 
| 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 305 | |> zero_var_indexes | 
| 50705 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
47954diff
changeset | 306 | |> new_skolem ? forall_intr_vars | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 307 | val (th, ctxt) = Variable.import true [th] ctxt |>> snd |>> the_single | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
52031diff
changeset | 308 | val th = th |> Conv.fconv_rule (Object_Logic.atomize ctxt) | 
| 59632 | 309 | |> cong_extensionalize_thm ctxt | 
| 47954 
aada9fd08b58
make higher-order goals more first-order via extensionality
 blanchet parents: 
47953diff
changeset | 310 | |> abs_extensionalize_thm ctxt | 
| 
aada9fd08b58
make higher-order goals more first-order via extensionality
 blanchet parents: 
47953diff
changeset | 311 | |> make_nnf ctxt | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 312 | in | 
| 50705 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
47954diff
changeset | 313 | if new_skolem then | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 314 | let | 
| 39901 
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
 blanchet parents: 
39900diff
changeset | 315 | fun skolemize choice_ths = | 
| 39950 | 316 | skolemize_with_choice_theorems ctxt choice_ths | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50705diff
changeset | 317 |           #> simplify (ss_only @{thms all_simps[symmetric]} ctxt)
 | 
| 42347 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 318 | val no_choice = null choice_ths | 
| 39901 
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
 blanchet parents: 
39900diff
changeset | 319 | val pull_out = | 
| 42347 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 320 | if no_choice then | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
50705diff
changeset | 321 |             simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]} ctxt)
 | 
| 42347 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 322 | else | 
| 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 323 | skolemize choice_ths | 
| 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 324 | val discharger_th = th |> pull_out | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 325 | val discharger_th = | 
| 59582 | 326 | discharger_th |> has_too_many_clauses ctxt (Thm.concl_of discharger_th) | 
| 42347 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 327 | ? (to_definitional_cnf_with_quantifiers ctxt | 
| 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 328 | #> pull_out) | 
| 42099 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 329 | val zapped_th = | 
| 59582 | 330 | discharger_th |> Thm.prop_of |> rename_bound_vars_to_be_zapped ax_no | 
| 40263 | 331 | |> (if no_choice then | 
| 59582 | 332 | Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> Thm.cprop_of | 
| 40263 | 333 | else | 
| 59632 | 334 | Thm.cterm_of ctxt) | 
| 42099 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 335 | |> zap true | 
| 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 336 | val fixes = | 
| 59582 | 337 | [] |> Term.add_free_names (Thm.prop_of zapped_th) | 
| 42335 | 338 | |> filter is_zapped_var_name | 
| 42269 | 339 | val ctxt' = ctxt |> Variable.add_fixes_direct fixes | 
| 42099 
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
 blanchet parents: 
42098diff
changeset | 340 | val fully_skolemized_t = | 
| 42333 | 341 | zapped_th |> singleton (Variable.export ctxt' ctxt) | 
| 59582 | 342 | |> Thm.cprop_of |> Thm.dest_equals |> snd |> Thm.term_of | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 343 | in | 
| 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 344 | if exists_subterm (fn Var ((s, _), _) => | 
| 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 345 | String.isPrefix new_skolem_var_prefix s | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 346 | | _ => false) fully_skolemized_t then | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 347 | let | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 348 | val (fully_skolemized_ct, ctxt) = | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 349 | Variable.import_terms true [fully_skolemized_t] ctxt | 
| 59632 | 350 | |>> the_single |>> Thm.cterm_of ctxt | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 351 | in | 
| 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 352 | (SOME (discharger_th, fully_skolemized_ct), | 
| 40262 
8403085384eb
ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
 blanchet parents: 
40261diff
changeset | 353 | (Thm.assume fully_skolemized_ct, ctxt)) | 
| 40260 
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
 blanchet parents: 
39962diff
changeset | 354 | end | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 355 | else | 
| 40262 
8403085384eb
ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
 blanchet parents: 
40261diff
changeset | 356 | (NONE, (th, ctxt)) | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 357 | end | 
| 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 358 | else | 
| 59582 | 359 | (NONE, (th |> has_too_many_clauses ctxt (Thm.concl_of th) | 
| 42347 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 360 | ? to_definitional_cnf_with_quantifiers ctxt, ctxt)) | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 361 | end | 
| 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 362 | |
| 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 363 | (* Convert a theorem to CNF, with additional premises due to skolemization. *) | 
| 50705 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
47954diff
changeset | 364 | fun cnf_axiom ctxt0 new_skolem combinators ax_no th = | 
| 37626 
1146291fe718
move blacklisting completely out of the clausifier;
 blanchet parents: 
37620diff
changeset | 365 | let | 
| 42361 | 366 | val thy = Proof_Context.theory_of ctxt0 | 
| 39950 | 367 | val choice_ths = choice_theorems thy | 
| 40262 
8403085384eb
ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
 blanchet parents: 
40261diff
changeset | 368 | val (opt, (nnf_th, ctxt)) = | 
| 50705 
0e943b33d907
use new skolemizer for reconstructing skolemization steps in Isar proofs (because the old skolemizer messes up the order of the Skolem arguments)
 blanchet parents: 
47954diff
changeset | 369 | nnf_axiom choice_ths new_skolem ax_no th ctxt0 | 
| 39894 
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
 blanchet parents: 
39891diff
changeset | 370 | fun clausify th = | 
| 59165 | 371 | make_cnf | 
| 372 | (if new_skolem orelse null choice_ths then [] | |
| 373 | else map (old_skolem_theorem_of_def ctxt) (old_skolem_defs th)) | |
| 374 | th ctxt | |
| 42347 
53e444ecb525
properly implemented definitional CNF for the new Skolemizer (and moved the code for the old Skolemizer -- tuning), removing big chunks for experimental/debugging code
 blanchet parents: 
42336diff
changeset | 375 | val (cnf_ths, ctxt) = clausify nnf_th | 
| 39894 
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
 blanchet parents: 
39891diff
changeset | 376 | fun intr_imp ct th = | 
| 60642 
48dd1cefb4ae
simplified Thm.instantiate and derivatives: the LHS refers to non-certified variables -- this merely serves as index into already certified structures (or is ignored);
 wenzelm parents: 
60328diff
changeset | 377 |       Thm.instantiate ([], [((("i", 0), @{typ nat}), Thm.cterm_of ctxt (HOLogic.mk_nat ax_no))])
 | 
| 39962 
d42ddd7407ca
qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
 blanchet parents: 
39950diff
changeset | 378 |                       (zero_var_indexes @{thm skolem_COMBK_D})
 | 
| 39894 
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
 blanchet parents: 
39891diff
changeset | 379 | RS Thm.implies_intr ct th | 
| 37626 
1146291fe718
move blacklisting completely out of the clausifier;
 blanchet parents: 
37620diff
changeset | 380 | in | 
| 39897 | 381 | (opt |> Option.map (I #>> singleton (Variable.export ctxt ctxt0) | 
| 59582 | 382 | ##> (Thm.term_of #> HOLogic.dest_Trueprop | 
| 39897 | 383 | #> singleton (Variable.export_terms ctxt ctxt0))), | 
| 55236 | 384 | cnf_ths |> map (combinators ? introduce_combinators_in_theorem ctxt | 
| 39894 
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
 blanchet parents: 
39891diff
changeset | 385 | #> (case opt of SOME (_, ct) => intr_imp ct | NONE => I)) | 
| 39897 | 386 | |> Variable.export ctxt ctxt0 | 
| 39950 | 387 | |> finish_cnf | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 388 | |> map Thm.close_derivation) | 
| 37626 
1146291fe718
move blacklisting completely out of the clausifier;
 blanchet parents: 
37620diff
changeset | 389 | end | 
| 39887 
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
 blanchet parents: 
39886diff
changeset | 390 | handle THM _ => (NONE, []) | 
| 27184 | 391 | |
| 20461 
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
 wenzelm parents: 
20457diff
changeset | 392 | end; |