author | wenzelm |
Tue, 22 Aug 2023 12:18:31 +0200 | |
changeset 78565 | 05de3e068312 |
parent 67613 | ce654b0e6d69 |
permissions | -rw-r--r-- |
43197 | 1 |
(* Title: HOL/Metis_Examples/Clausification.thy |
42338 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
||
43197 | 4 |
Example that exercises Metis's Clausifier. |
42338 | 5 |
*) |
6 |
||
63167 | 7 |
section \<open>Example that Exercises Metis's Clausifier\<close> |
43197 | 8 |
|
9 |
theory Clausification |
|
42338 | 10 |
imports Complex_Main |
11 |
begin |
|
12 |
||
42747
f132d13fcf75
use the same code for extensionalization in Metis and Sledgehammer and generalize that code so that it gracefully handles negations (e.g. negated conjecture), formulas of the form (%x. t) = u, etc.
blanchet
parents:
42350
diff
changeset
|
13 |
|
63167 | 14 |
text \<open>Definitional CNF for facts\<close> |
42345 | 15 |
|
16 |
declare [[meson_max_clauses = 10]] |
|
17 |
||
18 |
axiomatization q :: "nat \<Rightarrow> nat \<Rightarrow> bool" where |
|
19 |
qax: "\<exists>b. \<forall>a. (q b a \<and> q 0 0 \<and> q 1 a \<and> q a 1) \<or> (q 0 1 \<and> q 1 0 \<and> q a b \<and> q 1 1)" |
|
20 |
||
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:
50696
diff
changeset
|
21 |
declare [[metis_new_skolem = false]] |
42345 | 22 |
|
23 |
lemma "\<exists>b. \<forall>a. (q b a \<or> q a b)" |
|
24 |
by (metis qax) |
|
25 |
||
26 |
lemma "\<exists>b. \<forall>a. (q b a \<or> q a b)" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
27 |
by (metis (full_types) qax) |
42345 | 28 |
|
29 |
lemma "\<exists>b. \<forall>a. (q b a \<and> q 0 0 \<and> q 1 a \<and> q a 1) \<or> (q 0 1 \<and> q 1 0 \<and> q a b \<and> q 1 1)" |
|
30 |
by (metis qax) |
|
31 |
||
32 |
lemma "\<exists>b. \<forall>a. (q b a \<and> q 0 0 \<and> q 1 a \<and> q a 1) \<or> (q 0 1 \<and> q 1 0 \<and> q a b \<and> q 1 1)" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
33 |
by (metis (full_types) qax) |
42345 | 34 |
|
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:
50696
diff
changeset
|
35 |
declare [[metis_new_skolem]] |
42345 | 36 |
|
37 |
lemma "\<exists>b. \<forall>a. (q b a \<or> q a b)" |
|
38 |
by (metis qax) |
|
39 |
||
40 |
lemma "\<exists>b. \<forall>a. (q b a \<or> q a b)" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
41 |
by (metis (full_types) qax) |
42345 | 42 |
|
43 |
lemma "\<exists>b. \<forall>a. (q b a \<and> q 0 0 \<and> q 1 a \<and> q a 1) \<or> (q 0 1 \<and> q 1 0 \<and> q a b \<and> q 1 1)" |
|
44 |
by (metis qax) |
|
45 |
||
46 |
lemma "\<exists>b. \<forall>a. (q b a \<and> q 0 0 \<and> q 1 a \<and> q a 1) \<or> (q 0 1 \<and> q 1 0 \<and> q a b \<and> q 1 1)" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
47 |
by (metis (full_types) qax) |
42345 | 48 |
|
49 |
declare [[meson_max_clauses = 60]] |
|
50 |
||
51 |
axiomatization r :: "nat \<Rightarrow> nat \<Rightarrow> bool" where |
|
52 |
rax: "(r 0 0 \<and> r 0 1 \<and> r 0 2 \<and> r 0 3) \<or> |
|
53 |
(r 1 0 \<and> r 1 1 \<and> r 1 2 \<and> r 1 3) \<or> |
|
54 |
(r 2 0 \<and> r 2 1 \<and> r 2 2 \<and> r 2 3) \<or> |
|
55 |
(r 3 0 \<and> r 3 1 \<and> r 3 2 \<and> r 3 3)" |
|
56 |
||
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:
50696
diff
changeset
|
57 |
declare [[metis_new_skolem = false]] |
42345 | 58 |
|
59 |
lemma "r 0 0 \<or> r 1 1 \<or> r 2 2 \<or> r 3 3" |
|
60 |
by (metis rax) |
|
61 |
||
62 |
lemma "r 0 0 \<or> r 1 1 \<or> r 2 2 \<or> r 3 3" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
63 |
by (metis (full_types) rax) |
42345 | 64 |
|
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:
50696
diff
changeset
|
65 |
declare [[metis_new_skolem]] |
42345 | 66 |
|
67 |
lemma "r 0 0 \<or> r 1 1 \<or> r 2 2 \<or> r 3 3" |
|
68 |
by (metis rax) |
|
69 |
||
70 |
lemma "r 0 0 \<or> r 1 1 \<or> r 2 2 \<or> r 3 3" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
71 |
by (metis (full_types) rax) |
42345 | 72 |
|
42350 | 73 |
lemma "(r 0 0 \<and> r 0 1 \<and> r 0 2 \<and> r 0 3) \<or> |
74 |
(r 1 0 \<and> r 1 1 \<and> r 1 2 \<and> r 1 3) \<or> |
|
75 |
(r 2 0 \<and> r 2 1 \<and> r 2 2 \<and> r 2 3) \<or> |
|
76 |
(r 3 0 \<and> r 3 1 \<and> r 3 2 \<and> r 3 3)" |
|
77 |
by (metis rax) |
|
78 |
||
79 |
lemma "(r 0 0 \<and> r 0 1 \<and> r 0 2 \<and> r 0 3) \<or> |
|
80 |
(r 1 0 \<and> r 1 1 \<and> r 1 2 \<and> r 1 3) \<or> |
|
81 |
(r 2 0 \<and> r 2 1 \<and> r 2 2 \<and> r 2 3) \<or> |
|
82 |
(r 3 0 \<and> r 3 1 \<and> r 3 2 \<and> r 3 3)" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
83 |
by (metis (full_types) rax) |
42350 | 84 |
|
42747
f132d13fcf75
use the same code for extensionalization in Metis and Sledgehammer and generalize that code so that it gracefully handles negations (e.g. negated conjecture), formulas of the form (%x. t) = u, etc.
blanchet
parents:
42350
diff
changeset
|
85 |
|
63167 | 86 |
text \<open>Definitional CNF for goal\<close> |
42345 | 87 |
|
42338 | 88 |
axiomatization p :: "nat \<Rightarrow> nat \<Rightarrow> bool" where |
42345 | 89 |
pax: "\<exists>b. \<forall>a. (p b a \<and> p 0 0 \<and> p 1 a) \<or> (p 0 1 \<and> p 1 0 \<and> p a b)" |
42338 | 90 |
|
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:
50696
diff
changeset
|
91 |
declare [[metis_new_skolem = false]] |
42338 | 92 |
|
93 |
lemma "\<exists>b. \<forall>a. \<exists>x. (p b a \<or> x) \<and> (p 0 0 \<or> x) \<and> (p 1 a \<or> x) \<and> |
|
94 |
(p 0 1 \<or> \<not> x) \<and> (p 1 0 \<or> \<not> x) \<and> (p a b \<or> \<not> x)" |
|
95 |
by (metis pax) |
|
96 |
||
97 |
lemma "\<exists>b. \<forall>a. \<exists>x. (p b a \<or> x) \<and> (p 0 0 \<or> x) \<and> (p 1 a \<or> x) \<and> |
|
98 |
(p 0 1 \<or> \<not> x) \<and> (p 1 0 \<or> \<not> x) \<and> (p a b \<or> \<not> x)" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
99 |
by (metis (full_types) pax) |
42338 | 100 |
|
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:
50696
diff
changeset
|
101 |
declare [[metis_new_skolem]] |
42338 | 102 |
|
103 |
lemma "\<exists>b. \<forall>a. \<exists>x. (p b a \<or> x) \<and> (p 0 0 \<or> x) \<and> (p 1 a \<or> x) \<and> |
|
104 |
(p 0 1 \<or> \<not> x) \<and> (p 1 0 \<or> \<not> x) \<and> (p a b \<or> \<not> x)" |
|
105 |
by (metis pax) |
|
106 |
||
107 |
lemma "\<exists>b. \<forall>a. \<exists>x. (p b a \<or> x) \<and> (p 0 0 \<or> x) \<and> (p 1 a \<or> x) \<and> |
|
108 |
(p 0 1 \<or> \<not> x) \<and> (p 1 0 \<or> \<not> x) \<and> (p a b \<or> \<not> x)" |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
109 |
by (metis (full_types) pax) |
42338 | 110 |
|
42747
f132d13fcf75
use the same code for extensionalization in Metis and Sledgehammer and generalize that code so that it gracefully handles negations (e.g. negated conjecture), formulas of the form (%x. t) = u, etc.
blanchet
parents:
42350
diff
changeset
|
111 |
|
63167 | 112 |
text \<open>New Skolemizer\<close> |
42338 | 113 |
|
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:
50696
diff
changeset
|
114 |
declare [[metis_new_skolem]] |
42338 | 115 |
|
116 |
lemma |
|
117 |
fixes x :: real |
|
61969 | 118 |
assumes fn_le: "!!n. f n \<le> x" and 1: "f \<longlonglongrightarrow> lim f" |
42338 | 119 |
shows "lim f \<le> x" |
120 |
by (metis 1 LIMSEQ_le_const2 fn_le) |
|
121 |
||
122 |
definition |
|
123 |
bounded :: "'a::metric_space set \<Rightarrow> bool" where |
|
124 |
"bounded S \<longleftrightarrow> (\<exists>x eee. \<forall>y\<in>S. dist x y \<le> eee)" |
|
125 |
||
126 |
lemma "bounded T \<Longrightarrow> S \<subseteq> T ==> bounded S" |
|
127 |
by (metis bounded_def subset_eq) |
|
128 |
||
129 |
lemma |
|
47308 | 130 |
assumes a: "Quotient R Abs Rep T" |
42338 | 131 |
shows "symp R" |
132 |
using a unfolding Quotient_def using sympI |
|
43228
2ed2f092e990
obsoleted "metisFT", and added "no_types" version of Metis as fallback to Sledgehammer after noticing how useful it can be
blanchet
parents:
43197
diff
changeset
|
133 |
by (metis (full_types)) |
42338 | 134 |
|
135 |
lemma |
|
136 |
"(\<exists>x \<in> set xs. P x) \<longleftrightarrow> |
|
42342
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42340
diff
changeset
|
137 |
(\<exists>ys x zs. xs = ys @ x # zs \<and> P x \<and> (\<forall>z \<in> set zs. \<not> P z))" |
50696
85f944352d55
get rid of two-year-old hack, now that the "metis" skolemizer no longer gets stuck in HO unification
blanchet
parents:
47308
diff
changeset
|
138 |
by (metis split_list_last_prop[where P = P] in_set_conv_decomp) |
42338 | 139 |
|
67613 | 140 |
lemma ex_tl: "\<exists>ys. tl ys = xs" |
55417
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
50705
diff
changeset
|
141 |
using list.sel(3) by fast |
42342
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42340
diff
changeset
|
142 |
|
61076 | 143 |
lemma "(\<exists>ys::nat list. tl ys = xs) \<and> (\<exists>bs::int list. tl bs = as)" |
42342
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42340
diff
changeset
|
144 |
by (metis ex_tl) |
6babd86a54a4
handle case where the same Skolem name is given different types in different subgoals in the new Skolemizer (this can happen if several type-instances of the same fact are needed by Metis, cf. example in "Clausify.thy") -- the solution reintroduces old code removed in a6725f293377
blanchet
parents:
42340
diff
changeset
|
145 |
|
42338 | 146 |
end |