author | wenzelm |
Sat, 30 Nov 2024 13:40:57 +0100 | |
changeset 81511 | 8cbc8bc6f382 |
parent 81254 | d3c0734059ee |
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:
36228
diff
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 |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
77879
diff
changeset
|
10 |
type clause_options = {new_skolem : bool, combs : bool, refl : bool} |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
11 |
val new_skolem_var_prefix : string |
42098
f978caf60bbe
more robust handling of variables in new Skolemizer
blanchet
parents:
42072
diff
changeset
|
12 |
val new_nonskolem_var_prefix : string |
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
13 |
val is_zapped_var_name : string -> bool |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45508
diff
changeset
|
14 |
val is_quasi_lambda_free : term -> bool |
60328 | 15 |
val introduce_combinators_in_cterm : Proof.context -> cterm -> thm |
55236 | 16 |
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:
39931
diff
changeset
|
17 |
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:
50705
diff
changeset
|
18 |
val ss_only : thm list -> Proof.context -> Proof.context |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
77879
diff
changeset
|
19 |
val cnf_axiom : Meson.simp_options -> Proof.context -> clause_options -> int -> thm |
45508 | 20 |
-> (thm * term) option * thm list |
21505 | 21 |
end; |
19196
62ee8c10d796
Added functions to retrieve local and global atpset rules.
mengj
parents:
19175
diff
changeset
|
22 |
|
39890 | 23 |
structure Meson_Clausify : MESON_CLAUSIFY = |
15997 | 24 |
struct |
15347 | 25 |
|
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
77879
diff
changeset
|
26 |
type clause_options = {new_skolem : bool, combs : bool, refl : bool} |
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
77879
diff
changeset
|
27 |
|
42072
1492ee6b8085
avoid illegal variable names, which then yield "Not an identifier" errors in "rename_tac"
blanchet
parents:
41225
diff
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:
41225
diff
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:
41225
diff
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:
39886
diff
changeset
|
31 |
|
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
32 |
fun is_zapped_var_name s = |
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
33 |
exists (fn prefix => String.isPrefix prefix s) |
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
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:
42098
diff
changeset
|
35 |
|
15997 | 36 |
(**** Transformation of Elimination Rules into First-Order Formulas****) |
15347 | 37 |
|
67149 | 38 |
val cfalse = Thm.cterm_of \<^theory_context>\<open>HOL\<close> \<^term>\<open>False\<close>; |
39 |
val ctp_false = Thm.cterm_of \<^theory_context>\<open>HOL\<close> (HOLogic.mk_Trueprop \<^term>\<open>False\<close>); |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
40 |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
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:
38000
diff
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:
42747
diff
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:
38632
diff
changeset
|
44 |
"Sledgehammer_Util".) *) |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
45 |
fun transform_elim_theorem th = |
59582 | 46 |
(case Thm.concl_of th of (*conclusion variable*) |
74346 | 47 |
\<^Const_>\<open>Trueprop for \<open>Var (v as (_, \<^Type>\<open>bool\<close>))\<close>\<close> => |
77879 | 48 |
Thm.instantiate (TVars.empty, Vars.make1 (v, cfalse)) th |
74346 | 49 |
| Var (v as (_, \<^Type>\<open>prop\<close>)) => |
77879 | 50 |
Thm.instantiate (TVars.empty, Vars.make1 (v, ctp_false)) th |
59582 | 51 |
| _ => th) |
15997 | 52 |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
53 |
|
16009 | 54 |
(**** SKOLEMIZATION BY INFERENCE (lcp) ****) |
55 |
||
39886
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents:
39721
diff
changeset
|
56 |
fun mk_old_skolem_term_wrapper t = |
74346 | 57 |
let val T = fastype_of t |
58 |
in \<^Const>\<open>Meson.skolem T for t\<close> end |
|
37410
2bf7e6136047
adjusted the polymorphism handling of Skolem constants so that proof reconstruction doesn't fail in "equality_inf"
blanchet
parents:
37403
diff
changeset
|
59 |
|
77263
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
60 |
fun eta_expand_All_Ex_arg ((cst as (Const _)) $ Abs (s, T, t')) = |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
61 |
cst $ Abs (s, T, eta_expand_All_Ex_arg t') |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
62 |
| eta_expand_All_Ex_arg ((cst as (Const (cst_s, cst_T))) $ t') = |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
63 |
if member (op =) [\<^const_name>\<open>All\<close>, \<^const_name>\<open>Ex\<close>] cst_s then |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
64 |
cst $ Abs (Name.uu, domain_type (domain_type cst_T), |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
65 |
incr_boundvars 1 (eta_expand_All_Ex_arg t') $ Bound 0) |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
66 |
else |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
67 |
cst $ eta_expand_All_Ex_arg t' |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
68 |
| eta_expand_All_Ex_arg (Abs (s, T, t')) = Abs (s, T, eta_expand_All_Ex_arg t') |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
69 |
| eta_expand_All_Ex_arg (t1 $ t2) = |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
70 |
eta_expand_All_Ex_arg t1 $ eta_expand_All_Ex_arg t2 |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
71 |
| eta_expand_All_Ex_arg t = t |
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
72 |
|
39931 | 73 |
fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') |
77263
27be31d7ad88
careful eta-contraction in Metis to keep argument to All and Ex expanded
blanchet
parents:
74610
diff
changeset
|
74 |
| beta_eta_in_abs_body t = eta_expand_All_Ex_arg (Envir.beta_eta_contract t) |
37512
ff72d3ddc898
this looks like the most appropriate place to do the beta-eta-contraction
blanchet
parents:
37511
diff
changeset
|
75 |
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
76 |
(*Traverse a theorem, accumulating Skolem function definitions.*) |
39886
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents:
39721
diff
changeset
|
77 |
fun old_skolem_defs th = |
37399
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents:
37349
diff
changeset
|
78 |
let |
74346 | 79 |
fun dec_sko \<^Const_>\<open>Ex _ for \<open>body as Abs (_, T, p)\<close>\<close> rhss = |
37399
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents:
37349
diff
changeset
|
80 |
(*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:
37349
diff
changeset
|
81 |
let |
44121 | 82 |
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:
37498
diff
changeset
|
83 |
(* 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:
37498
diff
changeset
|
84 |
val rhs = |
44241 | 85 |
fold_rev (absfree o dest_Free) args |
86 |
(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:
39721
diff
changeset
|
87 |
|> 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:
37512
diff
changeset
|
88 |
val comb = list_comb (rhs, args) |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
89 |
in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74520
diff
changeset
|
90 |
| dec_sko \<^Const_>\<open>All _ for \<open>t as Abs _\<close>\<close> rhss = dec_sko (#2 (Term.dest_abs_global t)) rhss |
74346 | 91 |
| dec_sko \<^Const_>\<open>conj for p q\<close> rhss = rhss |> dec_sko p |> dec_sko q |
92 |
| dec_sko \<^Const_>\<open>disj for p q\<close> rhss = rhss |> dec_sko p |> dec_sko q |
|
93 |
| dec_sko \<^Const_>\<open>Trueprop for p\<close> rhss = dec_sko p rhss |
|
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
94 |
| dec_sko _ rhss = rhss |
59582 | 95 |
in dec_sko (Thm.prop_of th) [] end; |
20419 | 96 |
|
97 |
||
24827 | 98 |
(**** REPLACING ABSTRACTIONS BY COMBINATORS ****) |
20419 | 99 |
|
74346 | 100 |
fun is_quasi_lambda_free \<^Const_>\<open>Meson.skolem _ for _\<close> = true |
37416
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
blanchet
parents:
37410
diff
changeset
|
101 |
| is_quasi_lambda_free (t1 $ t2) = |
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
blanchet
parents:
37410
diff
changeset
|
102 |
is_quasi_lambda_free t1 andalso is_quasi_lambda_free t2 |
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
blanchet
parents:
37410
diff
changeset
|
103 |
| is_quasi_lambda_free (Abs _) = false |
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
blanchet
parents:
37410
diff
changeset
|
104 |
| is_quasi_lambda_free _ = true |
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
105 |
|
60328 | 106 |
fun abstract ctxt ct = |
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
107 |
let |
70157 | 108 |
val Abs (_, _, body) = Thm.term_of ct |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74520
diff
changeset
|
109 |
val (x, cbody) = Thm.dest_abs_global ct |
70159 | 110 |
val (A, cbodyT) = Thm.dest_funT (Thm.ctyp_of_cterm ct) |
70157 | 111 |
fun makeK () = Thm.instantiate' [SOME A, SOME cbodyT] [SOME cbody] @{thm abs_K} |
24827 | 112 |
in |
113 |
case body of |
|
114 |
Const _ => makeK() |
|
115 |
| Free _ => makeK() |
|
116 |
| Var _ => makeK() (*though Var isn't expected*) |
|
70157 | 117 |
| Bound 0 => Thm.instantiate' [SOME A] [] @{thm abs_I} (*identity: I*) |
24827 | 118 |
| rator$rand => |
42083
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents:
42072
diff
changeset
|
119 |
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:
42072
diff
changeset
|
120 |
if Term.is_dependent rand then (*S*) |
70157 | 121 |
let val crator = Thm.lambda x (Thm.dest_fun cbody) |
122 |
val crand = Thm.lambda x (Thm.dest_arg cbody) |
|
70159 | 123 |
val (C, B) = Thm.dest_funT (Thm.dest_ctyp1 (Thm.ctyp_of_cterm crator)) |
70157 | 124 |
val abs_S' = @{thm abs_S} |
125 |
|> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] |
|
59582 | 126 |
val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_S') |
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
127 |
in |
60328 | 128 |
Thm.transitive abs_S' (Conv.binop_conv (abstract ctxt) rhs) |
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
129 |
end |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
130 |
else (*C*) |
70157 | 131 |
let val crator = Thm.lambda x (Thm.dest_fun cbody) |
132 |
val crand = Thm.dest_arg cbody |
|
70159 | 133 |
val (C, B) = Thm.dest_funT (Thm.dest_ctyp1 (Thm.ctyp_of_cterm crator)) |
70157 | 134 |
val abs_C' = @{thm abs_C} |
135 |
|> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] |
|
59582 | 136 |
val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_C') |
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
137 |
in |
60328 | 138 |
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:
27174
diff
changeset
|
139 |
end |
42083
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents:
42072
diff
changeset
|
140 |
else if Term.is_dependent rand then (*B or eta*) |
36945 | 141 |
if rand = Bound 0 then Thm.eta_conversion ct |
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
142 |
else (*B*) |
70157 | 143 |
let val crator = Thm.dest_fun cbody |
144 |
val crand = Thm.lambda x (Thm.dest_arg cbody) |
|
70159 | 145 |
val (C, B) = Thm.dest_funT (Thm.ctyp_of_cterm crator) |
70157 | 146 |
val abs_B' = @{thm abs_B} |
147 |
|> instantiate'_normalize [SOME A, SOME B, SOME C] [SOME crator, SOME crand] |
|
59582 | 148 |
val (_,rhs) = Thm.dest_equals (Thm.cprop_of abs_B') |
60328 | 149 |
in Thm.transitive abs_B' (Conv.arg_conv (abstract ctxt) rhs) end |
150 |
else makeK () |
|
37349 | 151 |
| _ => raise Fail "abstract: Bad term" |
24827 | 152 |
end; |
20863
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
153 |
|
74520 | 154 |
(* Traverse a theorem, replacing lambda-abstractions with combinators. *) |
60328 | 155 |
fun introduce_combinators_in_cterm ctxt ct = |
59582 | 156 |
if is_quasi_lambda_free (Thm.term_of ct) then |
37349 | 157 |
Thm.reflexive ct |
59582 | 158 |
else case Thm.term_of ct of |
37349 | 159 |
Abs _ => |
160 |
let |
|
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74520
diff
changeset
|
161 |
val (cv, cta) = Thm.dest_abs_global ct |
59582 | 162 |
val (v, _) = dest_Free (Thm.term_of cv) |
60328 | 163 |
val u_th = introduce_combinators_in_cterm ctxt cta |
37349 | 164 |
val cu = Thm.rhs_of u_th |
60328 | 165 |
val comb_eq = abstract ctxt (Thm.lambda cv cu) |
37349 | 166 |
in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end |
167 |
| _ $ _ => |
|
168 |
let val (ct1, ct2) = Thm.dest_comb ct in |
|
60328 | 169 |
Thm.combination (introduce_combinators_in_cterm ctxt ct1) |
170 |
(introduce_combinators_in_cterm ctxt ct2) |
|
37349 | 171 |
end |
172 |
||
55236 | 173 |
fun introduce_combinators_in_theorem ctxt th = |
59582 | 174 |
if is_quasi_lambda_free (Thm.prop_of th) then |
37349 | 175 |
th |
24827 | 176 |
else |
37349 | 177 |
let |
178 |
val th = Drule.eta_contraction_rule th |
|
60328 | 179 |
val eqth = introduce_combinators_in_cterm ctxt (Thm.cprop_of th) |
37349 | 180 |
in Thm.equal_elim eqth th end |
181 |
handle THM (msg, _, _) => |
|
61268 | 182 |
(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:
55239
diff
changeset
|
183 |
"\nException message: " ^ msg); |
45511
9b0f8ca4388e
continued implementation of lambda-lifting in Metis
blanchet
parents:
45508
diff
changeset
|
184 |
(* A type variable of sort "{}" will make "abstraction" fail. *) |
37349 | 185 |
TrueI) |
16009 | 186 |
|
187 |
(*cterms are used throughout for efficiency*) |
|
69593 | 188 |
val cTrueprop = Thm.cterm_of \<^theory_context>\<open>HOL\<close> HOLogic.Trueprop; |
16009 | 189 |
|
190 |
(*Given an abstraction over n variables, replace the bound variables by free |
|
191 |
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:
20457
diff
changeset
|
192 |
fun c_variant_abs_multi (ct0, vars) = |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74520
diff
changeset
|
193 |
let val (cv,ct) = Thm.dest_abs_global ct0 |
16009 | 194 |
in c_variant_abs_multi (ct, cv::vars) end |
195 |
handle CTERM _ => (ct0, rev vars); |
|
196 |
||
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
197 |
(* 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:
37616
diff
changeset
|
198 |
an existential formula by a use of that function. |
74501 | 199 |
Example: "\<exists>x. x \<in> A \<and> x \<notin> B \<Longrightarrow> sko A B \<in> A \<and> sko A B \<notin> B" *) |
57466 | 200 |
fun old_skolem_theorem_of_def ctxt rhs0 = |
37399
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents:
37349
diff
changeset
|
201 |
let |
59632 | 202 |
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:
37616
diff
changeset
|
203 |
val rhs' = rhs |> Thm.dest_comb |> snd |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
204 |
val (ch, frees) = c_variant_abs_multi (rhs', []) |
59582 | 205 |
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:
37616
diff
changeset
|
206 |
val T = |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
207 |
case hilbert of |
67149 | 208 |
Const (_, Type (\<^type_name>\<open>fun\<close>, [_, T])) => T |
57466 | 209 |
| _ => raise TERM ("old_skolem_theorem_of_def: expected \"Eps\"", [hilbert]) |
59632 | 210 |
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:
46071
diff
changeset
|
211 |
val ex_tm = Thm.apply cTrueprop (Thm.apply cex cabs) |
37629 | 212 |
val conc = |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
213 |
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:
46071
diff
changeset
|
214 |
|> Drule.beta_conv cabs |> Thm.apply cTrueprop |
74530
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
215 |
fun tacf [prem] = |
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
216 |
rewrite_goals_tac ctxt @{thms skolem_def [abs_def]} |
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
217 |
THEN resolve_tac ctxt |
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
218 |
[(prem |> rewrite_rule ctxt @{thms skolem_def [abs_def]}) |
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
219 |
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:
37616
diff
changeset
|
220 |
in |
74530
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
wenzelm
parents:
74526
diff
changeset
|
221 |
Goal.prove_internal ctxt [ex_tm] conc tacf |
37629 | 222 |
|> forall_intr_list frees |
223 |
|> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) |
|
224 |
|> Thm.varifyT_global |
|
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
225 |
end |
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
226 |
|
42335 | 227 |
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:
38864
diff
changeset
|
228 |
let |
59582 | 229 |
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:
38864
diff
changeset
|
230 |
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:
38864
diff
changeset
|
231 |
val eqth = eqth RS @{thm TruepropI} |
dff91b90d74c
use definitional CNFs in Metis rather than plain CNF, following a suggestion by Joe Hurd;
blanchet
parents:
38864
diff
changeset
|
232 |
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:
38864
diff
changeset
|
233 |
|
39932
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents:
39931
diff
changeset
|
234 |
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:
39894
diff
changeset
|
235 |
(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:
39931
diff
changeset
|
236 |
"_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^ |
56811 | 237 |
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:
39894
diff
changeset
|
238 |
|
39899
608b108ec979
compute quantifier dependency graph in new skolemizer
blanchet
parents:
39897
diff
changeset
|
239 |
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:
39931
diff
changeset
|
240 |
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:
39931
diff
changeset
|
241 |
((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:
39931
diff
changeset
|
242 |
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:
39931
diff
changeset
|
243 |
end |
39897 | 244 |
|
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
245 |
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:
39962
diff
changeset
|
246 |
let |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
247 |
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:
39962
diff
changeset
|
248 |
case t of |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
249 |
(t1 as Const (s, _)) $ Abs (s', T, t') => |
67149 | 250 |
if s = \<^const_name>\<open>Pure.all\<close> orelse s = \<^const_name>\<open>All\<close> orelse |
251 |
s = \<^const_name>\<open>Ex\<close> then |
|
39932
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents:
39931
diff
changeset
|
252 |
let |
67149 | 253 |
val skolem = (pos = (s = \<^const_name>\<open>Ex\<close>)) |
39932
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents:
39931
diff
changeset
|
254 |
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:
39931
diff
changeset
|
255 |
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:
39931
diff
changeset
|
256 |
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:
39962
diff
changeset
|
257 |
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:
39962
diff
changeset
|
258 |
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:
39962
diff
changeset
|
259 |
else |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
260 |
t |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
261 |
| (t1 as Const (s, _)) $ t2 $ t3 => |
67149 | 262 |
if s = \<^const_name>\<open>Pure.imp\<close> orelse s = \<^const_name>\<open>HOL.implies\<close> then |
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
263 |
t1 $ aux cluster index_no (not pos) t2 $ aux cluster index_no pos t3 |
67149 | 264 |
else if s = \<^const_name>\<open>HOL.conj\<close> orelse |
265 |
s = \<^const_name>\<open>HOL.disj\<close> then |
|
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
266 |
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:
39962
diff
changeset
|
267 |
else |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
268 |
t |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
269 |
| (t1 as Const (s, _)) $ t2 => |
67149 | 270 |
if s = \<^const_name>\<open>Trueprop\<close> then |
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
271 |
t1 $ aux cluster index_no pos t2 |
67149 | 272 |
else if s = \<^const_name>\<open>Not\<close> then |
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
273 |
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:
39962
diff
changeset
|
274 |
else |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
275 |
t |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
276 |
| _ => t |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
277 |
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:
39962
diff
changeset
|
278 |
|
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
279 |
fun zap pos ct = |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
280 |
ct |
59582 | 281 |
|> (case Thm.term_of ct of |
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74520
diff
changeset
|
282 |
Const (s, _) $ Abs _ => |
67149 | 283 |
if s = \<^const_name>\<open>Pure.all\<close> orelse s = \<^const_name>\<open>All\<close> orelse |
284 |
s = \<^const_name>\<open>Ex\<close> then |
|
74525
c960bfcb91db
discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents:
74520
diff
changeset
|
285 |
Thm.dest_comb #> snd #> Thm.dest_abs_global #> snd #> zap pos |
39906 | 286 |
else |
287 |
Conv.all_conv |
|
288 |
| Const (s, _) $ _ $ _ => |
|
67149 | 289 |
if s = \<^const_name>\<open>Pure.imp\<close> orelse s = \<^const_name>\<open>implies\<close> then |
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
290 |
Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos) |
67149 | 291 |
else if s = \<^const_name>\<open>conj\<close> orelse s = \<^const_name>\<open>disj\<close> then |
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
292 |
Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos) |
39906 | 293 |
else |
294 |
Conv.all_conv |
|
295 |
| Const (s, _) $ _ => |
|
67149 | 296 |
if s = \<^const_name>\<open>Trueprop\<close> then Conv.arg_conv (zap pos) |
297 |
else if s = \<^const_name>\<open>Not\<close> then Conv.arg_conv (zap (not pos)) |
|
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
298 |
else Conv.all_conv |
39906 | 299 |
| _ => Conv.all_conv) |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
300 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
50705
diff
changeset
|
301 |
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:
39900
diff
changeset
|
302 |
|
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40260
diff
changeset
|
303 |
val cheat_choice = |
67149 | 304 |
\<^prop>\<open>\<forall>x. \<exists>y. Q x y \<Longrightarrow> \<exists>f. \<forall>x. Q x (f x)\<close> |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
305 |
|> Logic.varify_global |
69593 | 306 |
|> Skip_Proof.make_thm \<^theory> |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
307 |
|
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
308 |
(* Converts an Isabelle theorem into NNF. *) |
74051 | 309 |
fun nnf_axiom simp_options 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:
39886
diff
changeset
|
310 |
let |
42361 | 311 |
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:
39886
diff
changeset
|
312 |
val th = |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
313 |
th |> transform_elim_theorem |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
314 |
|> zero_var_indexes |
74245 | 315 |
|> new_skolem ? Thm.forall_intr_vars |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
316 |
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:
52031
diff
changeset
|
317 |
val th = th |> Conv.fconv_rule (Object_Logic.atomize ctxt) |
70507 | 318 |
|> Meson.cong_extensionalize_thm ctxt |
319 |
|> Meson.abs_extensionalize_thm ctxt |
|
74051 | 320 |
|> Meson.make_nnf simp_options ctxt |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
321 |
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:
47954
diff
changeset
|
322 |
if new_skolem then |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
323 |
let |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
324 |
fun skolemize choice_ths = |
74051 | 325 |
Meson.skolemize_with_choice_theorems simp_options ctxt choice_ths |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
50705
diff
changeset
|
326 |
#> 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:
42336
diff
changeset
|
327 |
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:
39900
diff
changeset
|
328 |
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:
42336
diff
changeset
|
329 |
if no_choice then |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
50705
diff
changeset
|
330 |
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:
42336
diff
changeset
|
331 |
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:
42336
diff
changeset
|
332 |
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:
42336
diff
changeset
|
333 |
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:
39962
diff
changeset
|
334 |
val discharger_th = |
70507 | 335 |
discharger_th |> Meson.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:
42336
diff
changeset
|
336 |
? (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:
42336
diff
changeset
|
337 |
#> pull_out) |
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
338 |
val zapped_th = |
59582 | 339 |
discharger_th |> Thm.prop_of |> rename_bound_vars_to_be_zapped ax_no |
40263 | 340 |
|> (if no_choice then |
59582 | 341 |
Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> Thm.cprop_of |
40263 | 342 |
else |
59632 | 343 |
Thm.cterm_of ctxt) |
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
344 |
|> zap true |
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
345 |
val fixes = |
59582 | 346 |
[] |> Term.add_free_names (Thm.prop_of zapped_th) |
42335 | 347 |
|> filter is_zapped_var_name |
42269 | 348 |
val ctxt' = ctxt |> Variable.add_fixes_direct fixes |
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
349 |
val fully_skolemized_t = |
42333 | 350 |
zapped_th |> singleton (Variable.export ctxt' ctxt) |
59582 | 351 |
|> 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:
39886
diff
changeset
|
352 |
in |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
353 |
if exists_subterm (fn Var ((s, _), _) => |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
354 |
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:
39962
diff
changeset
|
355 |
| _ => false) fully_skolemized_t then |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
356 |
let |
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
357 |
val (fully_skolemized_ct, ctxt) = |
70326 | 358 |
yield_singleton (Variable.import_terms true) fully_skolemized_t ctxt |
359 |
|>> Thm.cterm_of ctxt |
|
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
360 |
in |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
361 |
(SOME (discharger_th, fully_skolemized_ct), |
40262
8403085384eb
ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents:
40261
diff
changeset
|
362 |
(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:
39962
diff
changeset
|
363 |
end |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
364 |
else |
40262
8403085384eb
ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents:
40261
diff
changeset
|
365 |
(NONE, (th, ctxt)) |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
366 |
end |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
367 |
else |
70507 | 368 |
(NONE, (th |> Meson.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:
42336
diff
changeset
|
369 |
? 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:
39886
diff
changeset
|
370 |
end |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
371 |
|
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
372 |
(* Convert a theorem to CNF, with additional premises due to skolemization. *) |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
77879
diff
changeset
|
373 |
fun cnf_axiom simp_options ctxt0 {new_skolem, combs, refl} ax_no th = |
37626
1146291fe718
move blacklisting completely out of the clausifier;
blanchet
parents:
37620
diff
changeset
|
374 |
let |
70507 | 375 |
val choice_ths = Meson.choice_theorems (Proof_Context.theory_of ctxt0) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
70159
diff
changeset
|
376 |
val (opt, (nnf_th, ctxt1)) = |
74051 | 377 |
nnf_axiom simp_options choice_ths new_skolem ax_no th ctxt0 |
39894
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents:
39891
diff
changeset
|
378 |
fun clausify th = |
70507 | 379 |
Meson.make_cnf |
59165 | 380 |
(if new_skolem orelse null choice_ths then [] |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
70159
diff
changeset
|
381 |
else map (old_skolem_theorem_of_def ctxt1) (old_skolem_defs th)) |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
70159
diff
changeset
|
382 |
th ctxt1 |
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
70159
diff
changeset
|
383 |
val (cnf_ths, ctxt2) = clausify nnf_th |
39894
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents:
39891
diff
changeset
|
384 |
fun intr_imp ct th = |
74610 | 385 |
\<^instantiate>\<open>i = \<open>Thm.cterm_of ctxt2 (HOLogic.mk_nat ax_no)\<close> in |
386 |
lemma (schematic) \<open>skolem (COMBK P i) \<Longrightarrow> P\<close> for i :: nat |
|
387 |
by (rule iffD2 [OF skolem_COMBK_iff])\<close> |
|
39894
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents:
39891
diff
changeset
|
388 |
RS Thm.implies_intr ct th |
37626
1146291fe718
move blacklisting completely out of the clausifier;
blanchet
parents:
37620
diff
changeset
|
389 |
in |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
70159
diff
changeset
|
390 |
(opt |> Option.map (I #>> singleton (Variable.export ctxt2 ctxt0) |
59582 | 391 |
##> (Thm.term_of #> HOLogic.dest_Trueprop |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
70159
diff
changeset
|
392 |
#> singleton (Variable.export_terms ctxt2 ctxt0))), |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
77879
diff
changeset
|
393 |
cnf_ths |> map (combs ? introduce_combinators_in_theorem ctxt2 |
39894
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents:
39891
diff
changeset
|
394 |
#> (case opt of SOME (_, ct) => intr_imp ct | NONE => I)) |
70320
59258a3192bf
misc tuning and clarification, notably wrt. flow of context;
wenzelm
parents:
70159
diff
changeset
|
395 |
|> Variable.export ctxt2 ctxt0 |
81254
d3c0734059ee
variable instantiation in Sledgehammer and Metis
blanchet
parents:
77879
diff
changeset
|
396 |
|> Meson.finish_cnf refl |
70494 | 397 |
|> map (Thm.close_derivation \<^here>)) |
37626
1146291fe718
move blacklisting completely out of the clausifier;
blanchet
parents:
37620
diff
changeset
|
398 |
end |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
399 |
handle THM _ => (NONE, []) |
27184 | 400 |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
401 |
end; |