author | wenzelm |
Sat, 16 Apr 2011 18:11:20 +0200 | |
changeset 42364 | 8c674b3b8e44 |
parent 42361 | 23f352990944 |
child 42739 | 017e5dac8642 |
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 |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
10 |
val new_skolem_var_prefix : string |
42098
f978caf60bbe
more robust handling of variables in new Skolemizer
blanchet
parents:
42072
diff
changeset
|
11 |
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
|
12 |
val is_zapped_var_name : string -> bool |
38632
9cde57cdd0e3
treat "using X by metis" (more or less) the same as "by (metis X)"
blanchet
parents:
38610
diff
changeset
|
13 |
val extensionalize_theorem : thm -> thm |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
14 |
val introduce_combinators_in_cterm : cterm -> thm |
38028 | 15 |
val introduce_combinators_in_theorem : 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
|
16 |
val cluster_of_zapped_var_name : string -> (int * (int * int)) * bool |
42336
d63d43e85879
improve definitional CNF on goal by moving "not" past the quantifiers
blanchet
parents:
42335
diff
changeset
|
17 |
val ss_only : thm list -> simpset |
39897 | 18 |
val cnf_axiom : |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
19 |
Proof.context -> bool -> int -> thm -> (thm * term) option * thm list |
21505 | 20 |
end; |
19196
62ee8c10d796
Added functions to retrieve local and global atpset rules.
mengj
parents:
19175
diff
changeset
|
21 |
|
39890 | 22 |
structure Meson_Clausify : MESON_CLAUSIFY = |
15997 | 23 |
struct |
15347 | 24 |
|
39950 | 25 |
open Meson |
26 |
||
42072
1492ee6b8085
avoid illegal variable names, which then yield "Not an identifier" errors in "rename_tac"
blanchet
parents:
41225
diff
changeset
|
27 |
(* 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
|
28 |
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
|
29 |
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
|
30 |
|
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
31 |
fun is_zapped_var_name s = |
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
32 |
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
|
33 |
[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
|
34 |
|
15997 | 35 |
(**** Transformation of Elimination Rules into First-Order Formulas****) |
15347 | 36 |
|
29064 | 37 |
val cfalse = cterm_of @{theory HOL} HOLogic.false_const; |
38 |
val ctp_false = cterm_of @{theory HOL} (HOLogic.mk_Trueprop HOLogic.false_const); |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
39 |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
40 |
(* 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
|
41 |
predicate variable. Leaves other theorems unchanged. We simply instantiate |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
42 |
the conclusion variable to False. (Cf. "transform_elim_term" in |
38652
e063be321438
perform eta-expansion of quantifier bodies in Sledgehammer translation when needed + transform elim rules later;
blanchet
parents:
38632
diff
changeset
|
43 |
"Sledgehammer_Util".) *) |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
44 |
fun transform_elim_theorem th = |
21430
77651b6d9d6c
New transformation of eliminatino rules: we simply replace the final conclusion variable by False
paulson
parents:
21290
diff
changeset
|
45 |
case concl_of th of (*conclusion variable*) |
35963 | 46 |
@{const Trueprop} $ (v as Var (_, @{typ bool})) => |
29064 | 47 |
Thm.instantiate ([], [(cterm_of @{theory HOL} v, cfalse)]) th |
35963 | 48 |
| v as Var(_, @{typ prop}) => |
29064 | 49 |
Thm.instantiate ([], [(cterm_of @{theory HOL} v, ctp_false)]) th |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
50 |
| _ => th |
15997 | 51 |
|
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
52 |
|
16009 | 53 |
(**** SKOLEMIZATION BY INFERENCE (lcp) ****) |
54 |
||
39886
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents:
39721
diff
changeset
|
55 |
fun mk_old_skolem_term_wrapper t = |
37436 | 56 |
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:
39950
diff
changeset
|
57 |
Const (@{const_name Meson.skolem}, T --> T) $ t |
37436 | 58 |
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 |
|
39931 | 60 |
fun beta_eta_in_abs_body (Abs (s, T, t')) = Abs (s, T, beta_eta_in_abs_body t') |
61 |
| 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:
37511
diff
changeset
|
62 |
|
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
63 |
(*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
|
64 |
fun old_skolem_defs th = |
37399
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents:
37349
diff
changeset
|
65 |
let |
39376 | 66 |
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:
37349
diff
changeset
|
67 |
(*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
|
68 |
let |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
69 |
val args = OldTerm.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
|
70 |
(* 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
|
71 |
val rhs = |
7587b6e63454
thread original theorem along with CNF theorem, as a step toward killing the Skolem cache
blanchet
parents:
37498
diff
changeset
|
72 |
list_abs_free (map dest_Free args, |
39931 | 73 |
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
|
74 |
|> 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
|
75 |
val comb = list_comb (rhs, args) |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
76 |
in dec_sko (subst_bound (comb, p)) (rhs :: rhss) end |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
77 |
| 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:
37349
diff
changeset
|
78 |
(*Universal quant: insert a free variable into body and continue*) |
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents:
37349
diff
changeset
|
79 |
let val fname = Name.variant (OldTerm.add_term_names (p,[])) a |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
80 |
in dec_sko (subst_bound (Free(fname,T), p)) rhss end |
39906 | 81 |
| dec_sko (@{const conj} $ p $ q) rhss = rhss |> dec_sko p |> dec_sko q |
82 |
| 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:
37616
diff
changeset
|
83 |
| dec_sko (@{const Trueprop} $ p) rhss = dec_sko p rhss |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
84 |
| dec_sko _ rhss = rhss |
20419 | 85 |
in dec_sko (prop_of th) [] end; |
86 |
||
87 |
||
24827 | 88 |
(**** REPLACING ABSTRACTIONS BY COMBINATORS ****) |
20419 | 89 |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39268
diff
changeset
|
90 |
val fun_cong_all = @{thm fun_eq_iff [THEN iffD1]} |
20419 | 91 |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
92 |
(* Removes the lambdas from an equation of the form "t = (%x. u)". |
38608
01ed56c46259
beta eta contract the Sledgehammer conjecture (and also the axioms, although this might not be needed), just like Metis does (implicitly);
blanchet
parents:
38282
diff
changeset
|
93 |
(Cf. "extensionalize_term" in "Sledgehammer_Translate".) *) |
38000
c0b9efa8bfca
added extensionalization to Sledgehammer, mimicking what the Clausifier used to do
blanchet
parents:
37995
diff
changeset
|
94 |
fun extensionalize_theorem th = |
37540
48273d1ea331
better eta-expansion in Sledgehammer's clausification;
blanchet
parents:
37518
diff
changeset
|
95 |
case prop_of th of |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
96 |
_ $ (Const (@{const_name HOL.eq}, Type (_, [Type (@{type_name fun}, _), _])) |
39376 | 97 |
$ _ $ Abs _) => extensionalize_theorem (th RS fun_cong_all) |
37540
48273d1ea331
better eta-expansion in Sledgehammer's clausification;
blanchet
parents:
37518
diff
changeset
|
98 |
| _ => th |
20419 | 99 |
|
39962
d42ddd7407ca
qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
blanchet
parents:
39950
diff
changeset
|
100 |
fun is_quasi_lambda_free (Const (@{const_name Meson.skolem}, _) $ _) = 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 |
|
32010 | 106 |
val [f_B,g_B] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_B})); |
107 |
val [g_C,f_C] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_C})); |
|
108 |
val [f_S,g_S] = map (cterm_of @{theory}) (OldTerm.term_vars (prop_of @{thm abs_S})); |
|
20863
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
109 |
|
38282
319c59682c51
move Sledgehammer's HOL -> FOL translation to separate file (sledgehammer_translate.ML)
blanchet
parents:
38280
diff
changeset
|
110 |
(* FIXME: Requires more use of cterm constructors. *) |
24827 | 111 |
fun abstract ct = |
28544
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
112 |
let |
26743a1591f5
improved performance of skolem cache, due to parallel map;
wenzelm
parents:
28262
diff
changeset
|
113 |
val thy = theory_of_cterm ct |
25256
fe467fdf129a
Catch exceptions arising during the abstraction operation.
paulson
parents:
25243
diff
changeset
|
114 |
val Abs(x,_,body) = term_of ct |
35963 | 115 |
val Type(@{type_name fun}, [xT,bodyT]) = typ_of (ctyp_of_term ct) |
38005
b6555e9c5de4
prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents:
38001
diff
changeset
|
116 |
val cxT = ctyp_of thy xT |
b6555e9c5de4
prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents:
38001
diff
changeset
|
117 |
val cbodyT = ctyp_of thy bodyT |
b6555e9c5de4
prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents:
38001
diff
changeset
|
118 |
fun makeK () = |
b6555e9c5de4
prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents:
38001
diff
changeset
|
119 |
instantiate' [SOME cxT, SOME cbodyT] [SOME (cterm_of thy body)] |
b6555e9c5de4
prevent schematic variable clash in combinator-introduction code, when invoked from Sledgehammer (another consequence of the CNF -> FOF transition)
blanchet
parents:
38001
diff
changeset
|
120 |
@{thm abs_K} |
24827 | 121 |
in |
122 |
case body of |
|
123 |
Const _ => makeK() |
|
124 |
| Free _ => makeK() |
|
125 |
| Var _ => makeK() (*though Var isn't expected*) |
|
27184 | 126 |
| Bound 0 => instantiate' [SOME cxT] [] @{thm abs_I} (*identity: I*) |
24827 | 127 |
| 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
|
128 |
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
|
129 |
if Term.is_dependent rand then (*S*) |
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
130 |
let val crator = cterm_of thy (Abs(x,xT,rator)) |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
131 |
val crand = cterm_of thy (Abs(x,xT,rand)) |
27184 | 132 |
val abs_S' = cterm_instantiate [(f_S,crator),(g_S,crand)] @{thm abs_S} |
133 |
val (_,rhs) = Thm.dest_equals (cprop_of abs_S') |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
134 |
in |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
135 |
Thm.transitive abs_S' (Conv.binop_conv abstract rhs) |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
136 |
end |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
137 |
else (*C*) |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
138 |
let val crator = cterm_of thy (Abs(x,xT,rator)) |
27184 | 139 |
val abs_C' = cterm_instantiate [(f_C,crator),(g_C,cterm_of thy rand)] @{thm abs_C} |
140 |
val (_,rhs) = Thm.dest_equals (cprop_of abs_C') |
|
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
141 |
in |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
142 |
Thm.transitive abs_C' (Conv.fun_conv (Conv.arg_conv abstract) rhs) |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
143 |
end |
42083
e1209fc7ecdc
added Term.is_open and Term.is_dependent convenience, to cover common situations of loose bounds;
wenzelm
parents:
42072
diff
changeset
|
144 |
else if Term.is_dependent rand then (*B or eta*) |
36945 | 145 |
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
|
146 |
else (*B*) |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
147 |
let val crand = cterm_of thy (Abs(x,xT,rand)) |
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
148 |
val crator = cterm_of thy rator |
27184 | 149 |
val abs_B' = cterm_instantiate [(f_B,crator),(g_B,crand)] @{thm abs_B} |
150 |
val (_,rhs) = Thm.dest_equals (cprop_of abs_B') |
|
37349 | 151 |
in Thm.transitive abs_B' (Conv.arg_conv abstract rhs) end |
27179
8f29fed3dc9a
ResAxioms.cnf_axiom/cnf_rules_pairs: pass explicit theory context;
wenzelm
parents:
27174
diff
changeset
|
152 |
else makeK() |
37349 | 153 |
| _ => raise Fail "abstract: Bad term" |
24827 | 154 |
end; |
20863
4ee61dbf192d
improvements to abstraction, ensuring more re-use of abstraction functions
paulson
parents:
20789
diff
changeset
|
155 |
|
37349 | 156 |
(* Traverse a theorem, remplacing lambda-abstractions with combinators. *) |
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
157 |
fun introduce_combinators_in_cterm ct = |
37416
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
blanchet
parents:
37410
diff
changeset
|
158 |
if is_quasi_lambda_free (term_of ct) then |
37349 | 159 |
Thm.reflexive ct |
160 |
else case term_of ct of |
|
161 |
Abs _ => |
|
162 |
let |
|
163 |
val (cv, cta) = Thm.dest_abs NONE ct |
|
164 |
val (v, _) = dest_Free (term_of cv) |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
165 |
val u_th = introduce_combinators_in_cterm cta |
37349 | 166 |
val cu = Thm.rhs_of u_th |
167 |
val comb_eq = abstract (Thm.cabs cv cu) |
|
168 |
in Thm.transitive (Thm.abstract_rule v cv u_th) comb_eq end |
|
169 |
| _ $ _ => |
|
170 |
let val (ct1, ct2) = Thm.dest_comb ct in |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
171 |
Thm.combination (introduce_combinators_in_cterm ct1) |
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
172 |
(introduce_combinators_in_cterm ct2) |
37349 | 173 |
end |
174 |
||
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
175 |
fun introduce_combinators_in_theorem th = |
37416
d5ac8280497e
no point in introducing combinators for inlined Skolem functions
blanchet
parents:
37410
diff
changeset
|
176 |
if is_quasi_lambda_free (prop_of th) then |
37349 | 177 |
th |
24827 | 178 |
else |
37349 | 179 |
let |
180 |
val th = Drule.eta_contraction_rule th |
|
38001
a9b47b85ca24
reintroduced more preprocessing steps to Sledgehammer, adapted to the new FOF setting
blanchet
parents:
38000
diff
changeset
|
181 |
val eqth = introduce_combinators_in_cterm (cprop_of th) |
37349 | 182 |
in Thm.equal_elim eqth th end |
183 |
handle THM (msg, _, _) => |
|
184 |
(warning ("Error in the combinator translation of " ^ |
|
185 |
Display.string_of_thm_without_context th ^ |
|
186 |
"\nException message: " ^ msg ^ "."); |
|
187 |
(* A type variable of sort "{}" will make abstraction fail. *) |
|
188 |
TrueI) |
|
16009 | 189 |
|
190 |
(*cterms are used throughout for efficiency*) |
|
38280 | 191 |
val cTrueprop = cterm_of @{theory HOL} HOLogic.Trueprop; |
16009 | 192 |
|
193 |
(*Given an abstraction over n variables, replace the bound variables by free |
|
194 |
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
|
195 |
fun c_variant_abs_multi (ct0, vars) = |
16009 | 196 |
let val (cv,ct) = Thm.dest_abs NONE ct0 |
197 |
in c_variant_abs_multi (ct, cv::vars) end |
|
198 |
handle CTERM _ => (ct0, rev vars); |
|
199 |
||
39355 | 200 |
val skolem_def_raw = @{thms skolem_def_raw} |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
201 |
|
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
202 |
(* 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
|
203 |
an existential formula by a use of that function. |
18141
89e2e8bed08f
Skolemization by inference, but not quite finished
paulson
parents:
18009
diff
changeset
|
204 |
Example: "EX x. x : A & x ~: B ==> sko A B : A & sko A B ~: B" [.] *) |
39886
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents:
39721
diff
changeset
|
205 |
fun old_skolem_theorem_from_def thy rhs0 = |
37399
34f080a12063
proper polymorphic Skolemization of uncached facts + synchronization of caching and relevance filter
blanchet
parents:
37349
diff
changeset
|
206 |
let |
38280 | 207 |
val rhs = rhs0 |> Type.legacy_freeze_thaw |> #1 |> cterm_of thy |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
208 |
val rhs' = rhs |> Thm.dest_comb |> snd |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
209 |
val (ch, frees) = c_variant_abs_multi (rhs', []) |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
210 |
val (hilbert, cabs) = ch |> Thm.dest_comb |>> term_of |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
211 |
val T = |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
212 |
case hilbert of |
39941 | 213 |
Const (_, Type (@{type_name fun}, [_, T])) => T |
39886
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents:
39721
diff
changeset
|
214 |
| _ => raise TERM ("old_skolem_theorem_from_def: expected \"Eps\"", |
8a9f0c97d550
first step towards a new skolemizer that doesn't require "Eps"
blanchet
parents:
39721
diff
changeset
|
215 |
[hilbert]) |
38280 | 216 |
val cex = cterm_of thy (HOLogic.exists_const T) |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
217 |
val ex_tm = Thm.capply cTrueprop (Thm.capply cex cabs) |
37629 | 218 |
val conc = |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
219 |
Drule.list_comb (rhs, frees) |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
220 |
|> Drule.beta_conv cabs |> Thm.capply cTrueprop |
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
221 |
fun tacf [prem] = |
39355 | 222 |
rewrite_goals_tac skolem_def_raw |
39941 | 223 |
THEN rtac ((prem |> rewrite_rule skolem_def_raw) |
39949 | 224 |
RS Global_Theory.get_thm thy "Hilbert_Choice.someI_ex") 1 |
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
225 |
in |
37629 | 226 |
Goal.prove_internal [ex_tm] conc tacf |
227 |
|> forall_intr_list frees |
|
228 |
|> Thm.forall_elim_vars 0 (*Introduce Vars, but don't discharge defs.*) |
|
229 |
|> Thm.varifyT_global |
|
37617
f73cd4069f69
always perform "inline" skolemization, polymorphism or not, Skolem cache or not
blanchet
parents:
37616
diff
changeset
|
230 |
end |
24742
73b8b42a36b6
removal of some "ref"s from res_axioms.ML; a side-effect is that the ordering
paulson
parents:
24669
diff
changeset
|
231 |
|
42335 | 232 |
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
|
233 |
let |
42335 | 234 |
val eqth = cnf.make_cnfx_thm ctxt (HOLogic.dest_Trueprop (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
|
235 |
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
|
236 |
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
|
237 |
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
|
238 |
|
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
|
239 |
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
|
240 |
(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
|
241 |
"_" ^ string_of_int ax_no ^ "_" ^ string_of_int cluster_no ^ "_" ^ |
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40260
diff
changeset
|
242 |
string_of_int index_no ^ "_" ^ Name.desymbolize 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
|
243 |
|
39899
608b108ec979
compute quantifier dependency graph in new skolemizer
blanchet
parents:
39897
diff
changeset
|
244 |
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
|
245 |
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
|
246 |
((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
|
247 |
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
|
248 |
end |
39897 | 249 |
|
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
250 |
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
|
251 |
let |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
252 |
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
|
253 |
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
|
254 |
(t1 as Const (s, _)) $ Abs (s', T, t') => |
39906 | 255 |
if s = @{const_name all} orelse s = @{const_name All} orelse |
256 |
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:
39931
diff
changeset
|
257 |
let |
acde1b606b0e
reverted 0bfaaa81fc62, since "Thm.rename_boundvars" can't be relied upon to give unique names to bound variables
blanchet
parents:
39931
diff
changeset
|
258 |
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:
39931
diff
changeset
|
259 |
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
|
260 |
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
|
261 |
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
|
262 |
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
|
263 |
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
|
264 |
else |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
265 |
t |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
266 |
| (t1 as Const (s, _)) $ t2 $ t3 => |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
267 |
if s = @{const_name "==>"} orelse s = @{const_name HOL.implies} then |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
268 |
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:
39962
diff
changeset
|
269 |
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:
39962
diff
changeset
|
270 |
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:
39962
diff
changeset
|
271 |
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
|
272 |
else |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
273 |
t |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
274 |
| (t1 as Const (s, _)) $ t2 => |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
275 |
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:
39962
diff
changeset
|
276 |
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:
39962
diff
changeset
|
277 |
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:
39962
diff
changeset
|
278 |
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
|
279 |
else |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
280 |
t |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
281 |
| _ => t |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
282 |
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
|
283 |
|
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
284 |
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
|
285 |
ct |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
286 |
|> (case term_of ct of |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
287 |
Const (s, _) $ Abs (s', _, _) => |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
288 |
if s = @{const_name all} orelse s = @{const_name All} orelse |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
289 |
s = @{const_name Ex} then |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
290 |
Thm.dest_comb #> snd #> Thm.dest_abs (SOME s') #> snd #> zap pos |
39906 | 291 |
else |
292 |
Conv.all_conv |
|
293 |
| Const (s, _) $ _ $ _ => |
|
294 |
if s = @{const_name "==>"} 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:
39962
diff
changeset
|
295 |
Conv.combination_conv (Conv.arg_conv (zap (not pos))) (zap pos) |
39906 | 296 |
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:
39962
diff
changeset
|
297 |
Conv.combination_conv (Conv.arg_conv (zap pos)) (zap pos) |
39906 | 298 |
else |
299 |
Conv.all_conv |
|
300 |
| Const (s, _) $ _ => |
|
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
301 |
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:
39962
diff
changeset
|
302 |
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:
39962
diff
changeset
|
303 |
else Conv.all_conv |
39906 | 304 |
| _ => Conv.all_conv) |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
305 |
|
41225 | 306 |
fun ss_only ths = Simplifier.clear_ss HOL_basic_ss addsimps ths |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
307 |
|
40261
7a02144874f3
more work on new Skolemizer without Hilbert_Choice
blanchet
parents:
40260
diff
changeset
|
308 |
val cheat_choice = |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
309 |
@{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:
39900
diff
changeset
|
310 |
|> Logic.varify_global |
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
311 |
|> 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
|
312 |
|
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
313 |
(* Converts an Isabelle theorem into NNF. *) |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
314 |
fun nnf_axiom choice_ths new_skolemizer 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
|
315 |
let |
42361 | 316 |
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
|
317 |
val th = |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
318 |
th |> transform_elim_theorem |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
319 |
|> zero_var_indexes |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
320 |
|> new_skolemizer ? forall_intr_vars |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
321 |
val (th, ctxt) = Variable.import true [th] ctxt |>> snd |>> the_single |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
322 |
val th = th |> Conv.fconv_rule Object_Logic.atomize |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
323 |
|> extensionalize_theorem |
39950 | 324 |
|> make_nnf ctxt |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
325 |
in |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
326 |
if new_skolemizer then |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
327 |
let |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
328 |
fun skolemize choice_ths = |
39950 | 329 |
skolemize_with_choice_theorems ctxt choice_ths |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
330 |
#> simplify (ss_only @{thms all_simps[symmetric]}) |
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 |
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
|
332 |
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
|
333 |
if no_choice then |
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
|
334 |
simplify (ss_only @{thms all_simps[symmetric] ex_simps[symmetric]}) |
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
|
335 |
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
|
336 |
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
|
337 |
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
|
338 |
val 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
|
339 |
discharger_th |> has_too_many_clauses ctxt (concl_of discharger_th) |
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
|
340 |
? (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
|
341 |
#> pull_out) |
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
342 |
val zapped_th = |
40263 | 343 |
discharger_th |> prop_of |> rename_bound_vars_to_be_zapped ax_no |
344 |
|> (if no_choice then |
|
345 |
Skip_Proof.make_thm thy #> skolemize [cheat_choice] #> cprop_of |
|
346 |
else |
|
347 |
cterm_of thy) |
|
42099
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
348 |
|> zap true |
447fa058ab22
avoid evil "export_without_context", which breaks if there are local "fixes"
blanchet
parents:
42098
diff
changeset
|
349 |
val fixes = |
42335 | 350 |
[] |> Term.add_free_names (prop_of zapped_th) |
351 |
|> filter is_zapped_var_name |
|
42269 | 352 |
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
|
353 |
val fully_skolemized_t = |
42333 | 354 |
zapped_th |> singleton (Variable.export ctxt' ctxt) |
355 |
|> cprop_of |> Thm.dest_equals |> snd |> term_of |
|
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
356 |
in |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
357 |
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
|
358 |
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
|
359 |
| _ => 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
|
360 |
let |
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
361 |
val (fully_skolemized_ct, ctxt) = |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
362 |
Variable.import_terms true [fully_skolemized_t] ctxt |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
363 |
|>> the_single |>> cterm_of thy |
40260
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
364 |
in |
73ecbe2d432b
fix cluster numbering in the absense of Hilbert_Choice (reverts acde1b606b0e, effectively reintroducing most of 0bfaaa81fc62)
blanchet
parents:
39962
diff
changeset
|
365 |
(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
|
366 |
(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
|
367 |
end |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
368 |
else |
40262
8403085384eb
ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents:
40261
diff
changeset
|
369 |
(NONE, (th, 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 |
else |
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
|
372 |
(NONE, (th |> has_too_many_clauses ctxt (concl_of th) |
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
|
373 |
? 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
|
374 |
end |
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
375 |
|
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
376 |
(* Convert a theorem to CNF, with additional premises due to skolemization. *) |
39901
75d792edf634
make "cnf_axiom" work (after a fashion) in the absence of the axiom of choice
blanchet
parents:
39900
diff
changeset
|
377 |
fun cnf_axiom ctxt0 new_skolemizer ax_no th = |
37626
1146291fe718
move blacklisting completely out of the clausifier;
blanchet
parents:
37620
diff
changeset
|
378 |
let |
42361 | 379 |
val thy = Proof_Context.theory_of ctxt0 |
39950 | 380 |
val choice_ths = choice_theorems thy |
40262
8403085384eb
ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents:
40261
diff
changeset
|
381 |
val (opt, (nnf_th, ctxt)) = |
8403085384eb
ensure that MESON correctly preserves the name of variables (needed by the new Skolemizer)
blanchet
parents:
40261
diff
changeset
|
382 |
nnf_axiom choice_ths new_skolemizer ax_no th ctxt0 |
39894
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents:
39891
diff
changeset
|
383 |
fun clausify 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
|
384 |
make_cnf (if new_skolemizer orelse null choice_ths then [] |
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
|
385 |
else map (old_skolem_theorem_from_def thy) (old_skolem_defs th)) |
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
|
386 |
th 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
|
387 |
val (cnf_ths, ctxt) = clausify nnf_th |
39894
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents:
39891
diff
changeset
|
388 |
fun intr_imp ct th = |
39950 | 389 |
Thm.instantiate ([], map (pairself (cterm_of thy)) |
39962
d42ddd7407ca
qualify "Meson." and "Metis." constants so that the ML files can be reloaded later (e.g. for testing/development)
blanchet
parents:
39950
diff
changeset
|
390 |
[(Var (("i", 0), @{typ nat}), |
39902 | 391 |
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:
39950
diff
changeset
|
392 |
(zero_var_indexes @{thm skolem_COMBK_D}) |
39894
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents:
39891
diff
changeset
|
393 |
RS Thm.implies_intr ct th |
37626
1146291fe718
move blacklisting completely out of the clausifier;
blanchet
parents:
37620
diff
changeset
|
394 |
in |
39897 | 395 |
(opt |> Option.map (I #>> singleton (Variable.export ctxt ctxt0) |
396 |
##> (term_of #> HOLogic.dest_Trueprop |
|
397 |
#> singleton (Variable.export_terms ctxt ctxt0))), |
|
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
398 |
cnf_ths |> map (introduce_combinators_in_theorem |
39894
35ae5cf8c96a
encode number of skolem assumptions in them, for more efficient retrieval later
blanchet
parents:
39891
diff
changeset
|
399 |
#> (case opt of SOME (_, ct) => intr_imp ct | NONE => I)) |
39897 | 400 |
|> Variable.export ctxt ctxt0 |
39950 | 401 |
|> finish_cnf |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
402 |
|> map Thm.close_derivation) |
37626
1146291fe718
move blacklisting completely out of the clausifier;
blanchet
parents:
37620
diff
changeset
|
403 |
end |
39887
74939e2afb95
second step in introducing the new Skolemizer -- notably, added procedure for discharging Skolem assumptions
blanchet
parents:
39886
diff
changeset
|
404 |
handle THM _ => (NONE, []) |
27184 | 405 |
|
20461
d689ad772b2c
skolem_cache_thm: Drule.close_derivation on clauses preserves some space;
wenzelm
parents:
20457
diff
changeset
|
406 |
end; |