author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 67613 | ce654b0e6d69 |
permissions | -rw-r--r-- |
59667
651ea265d568
Removal of the file HOL/Number_Theory/Binomial!! And class field_char_0 now declared in Int.thy
paulson <lp15@cam.ac.uk>
parents:
55159
diff
changeset
|
1 |
theory Examples imports Complex_Main begin |
10294 | 2 |
|
39128
93a7365fb4ee
turned eta_contract into proper configuration option;
wenzelm
parents:
38798
diff
changeset
|
3 |
declare [[eta_contract = false]] |
10294 | 4 |
|
67406 | 5 |
text\<open>membership, intersection\<close> |
6 |
text\<open>difference and empty set\<close> |
|
7 |
text\<open>complement, union and universal set\<close> |
|
10294 | 8 |
|
9 |
lemma "(x \<in> A \<inter> B) = (x \<in> A \<and> x \<in> B)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
10 |
by blast |
10294 | 11 |
|
67406 | 12 |
text\<open> |
10294 | 13 |
@{thm[display] IntI[no_vars]} |
14 |
\rulename{IntI} |
|
15 |
||
16 |
@{thm[display] IntD1[no_vars]} |
|
17 |
\rulename{IntD1} |
|
18 |
||
19 |
@{thm[display] IntD2[no_vars]} |
|
20 |
\rulename{IntD2} |
|
67406 | 21 |
\<close> |
10294 | 22 |
|
23 |
lemma "(x \<in> -A) = (x \<notin> A)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
24 |
by blast |
10294 | 25 |
|
67406 | 26 |
text\<open> |
10294 | 27 |
@{thm[display] Compl_iff[no_vars]} |
28 |
\rulename{Compl_iff} |
|
67406 | 29 |
\<close> |
10294 | 30 |
|
31 |
lemma "- (A \<union> B) = -A \<inter> -B" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
32 |
by blast |
10294 | 33 |
|
67406 | 34 |
text\<open> |
10294 | 35 |
@{thm[display] Compl_Un[no_vars]} |
36 |
\rulename{Compl_Un} |
|
67406 | 37 |
\<close> |
10294 | 38 |
|
39 |
lemma "A-A = {}" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
40 |
by blast |
10294 | 41 |
|
67406 | 42 |
text\<open> |
10294 | 43 |
@{thm[display] Diff_disjoint[no_vars]} |
44 |
\rulename{Diff_disjoint} |
|
67406 | 45 |
\<close> |
10294 | 46 |
|
55159
608c157d743d
Replacing the theory Library/Binomial by Number_Theory/Binomial
paulson <lp15@cam.ac.uk>
parents:
48985
diff
changeset
|
47 |
|
10294 | 48 |
|
49 |
lemma "A \<union> -A = UNIV" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
50 |
by blast |
10294 | 51 |
|
67406 | 52 |
text\<open> |
10294 | 53 |
@{thm[display] Compl_partition[no_vars]} |
54 |
\rulename{Compl_partition} |
|
67406 | 55 |
\<close> |
10294 | 56 |
|
67406 | 57 |
text\<open>subset relation\<close> |
10294 | 58 |
|
59 |
||
67406 | 60 |
text\<open> |
10294 | 61 |
@{thm[display] subsetI[no_vars]} |
62 |
\rulename{subsetI} |
|
63 |
||
64 |
@{thm[display] subsetD[no_vars]} |
|
65 |
\rulename{subsetD} |
|
67406 | 66 |
\<close> |
10294 | 67 |
|
68 |
lemma "((A \<union> B) \<subseteq> C) = (A \<subseteq> C \<and> B \<subseteq> C)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
69 |
by blast |
10294 | 70 |
|
67406 | 71 |
text\<open> |
10294 | 72 |
@{thm[display] Un_subset_iff[no_vars]} |
73 |
\rulename{Un_subset_iff} |
|
67406 | 74 |
\<close> |
10294 | 75 |
|
76 |
lemma "(A \<subseteq> -B) = (B \<subseteq> -A)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
77 |
by blast |
10294 | 78 |
|
79 |
lemma "(A <= -B) = (B <= -A)" |
|
80 |
oops |
|
81 |
||
67406 | 82 |
text\<open>ASCII version: blast fails because of overloading because |
83 |
it doesn't have to be sets\<close> |
|
10294 | 84 |
|
85 |
lemma "((A:: 'a set) <= -B) = (B <= -A)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
86 |
by blast |
10294 | 87 |
|
67406 | 88 |
text\<open>A type constraint lets it work\<close> |
10294 | 89 |
|
67406 | 90 |
text\<open>An issue here: how do we discuss the distinction between ASCII and |
91 |
symbol notation? Here the latter disambiguates.\<close> |
|
10294 | 92 |
|
93 |
||
67406 | 94 |
text\<open> |
10294 | 95 |
set extensionality |
96 |
||
39795 | 97 |
@{thm[display] set_eqI[no_vars]} |
98 |
\rulename{set_eqI} |
|
10294 | 99 |
|
100 |
@{thm[display] equalityI[no_vars]} |
|
101 |
\rulename{equalityI} |
|
102 |
||
103 |
@{thm[display] equalityE[no_vars]} |
|
104 |
\rulename{equalityE} |
|
67406 | 105 |
\<close> |
10294 | 106 |
|
107 |
||
67406 | 108 |
text\<open>finite sets: insertion and membership relation\<close> |
109 |
text\<open>finite set notation\<close> |
|
10294 | 110 |
|
111 |
lemma "insert x A = {x} \<union> A" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
112 |
by blast |
10294 | 113 |
|
67406 | 114 |
text\<open> |
10294 | 115 |
@{thm[display] insert_is_Un[no_vars]} |
116 |
\rulename{insert_is_Un} |
|
67406 | 117 |
\<close> |
10294 | 118 |
|
119 |
lemma "{a,b} \<union> {c,d} = {a,b,c,d}" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
120 |
by blast |
10294 | 121 |
|
122 |
lemma "{a,b} \<inter> {b,c} = {b}" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
123 |
apply auto |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
124 |
oops |
67406 | 125 |
text\<open>fails because it isn't valid\<close> |
10294 | 126 |
|
127 |
lemma "{a,b} \<inter> {b,c} = (if a=c then {a,b} else {b})" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
128 |
apply simp |
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
129 |
by blast |
10294 | 130 |
|
67406 | 131 |
text\<open>or just force or auto. blast alone can't handle the if-then-else\<close> |
10294 | 132 |
|
67406 | 133 |
text\<open>next: some comprehension examples\<close> |
10294 | 134 |
|
135 |
lemma "(a \<in> {z. P z}) = P a" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
136 |
by blast |
10294 | 137 |
|
67406 | 138 |
text\<open> |
10294 | 139 |
@{thm[display] mem_Collect_eq[no_vars]} |
140 |
\rulename{mem_Collect_eq} |
|
67406 | 141 |
\<close> |
10294 | 142 |
|
143 |
lemma "{x. x \<in> A} = A" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
144 |
by blast |
55159
608c157d743d
Replacing the theory Library/Binomial by Number_Theory/Binomial
paulson <lp15@cam.ac.uk>
parents:
48985
diff
changeset
|
145 |
|
67406 | 146 |
text\<open> |
10294 | 147 |
@{thm[display] Collect_mem_eq[no_vars]} |
148 |
\rulename{Collect_mem_eq} |
|
67406 | 149 |
\<close> |
10294 | 150 |
|
151 |
lemma "{x. P x \<or> x \<in> A} = {x. P x} \<union> A" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
152 |
by blast |
10294 | 153 |
|
154 |
lemma "{x. P x \<longrightarrow> Q x} = -{x. P x} \<union> {x. Q x}" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
155 |
by blast |
10294 | 156 |
|
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
32833
diff
changeset
|
157 |
definition prime :: "nat set" where |
67613 | 158 |
"prime == {p. 1<p & (\<forall>m. m dvd p \<longrightarrow> m=1 \<or> m=p)}" |
10294 | 159 |
|
55159
608c157d743d
Replacing the theory Library/Binomial by Number_Theory/Binomial
paulson <lp15@cam.ac.uk>
parents:
48985
diff
changeset
|
160 |
lemma "{p*q | p q. p\<in>prime \<and> q\<in>prime} = |
10294 | 161 |
{z. \<exists>p q. z = p*q \<and> p\<in>prime \<and> q\<in>prime}" |
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
162 |
by (rule refl) |
10294 | 163 |
|
67406 | 164 |
text\<open>binders\<close> |
10294 | 165 |
|
67406 | 166 |
text\<open>bounded quantifiers\<close> |
10294 | 167 |
|
168 |
lemma "(\<exists>x\<in>A. P x) = (\<exists>x. x\<in>A \<and> P x)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
169 |
by blast |
10294 | 170 |
|
67406 | 171 |
text\<open> |
10294 | 172 |
@{thm[display] bexI[no_vars]} |
173 |
\rulename{bexI} |
|
67406 | 174 |
\<close> |
10294 | 175 |
|
67406 | 176 |
text\<open> |
10294 | 177 |
@{thm[display] bexE[no_vars]} |
178 |
\rulename{bexE} |
|
67406 | 179 |
\<close> |
10294 | 180 |
|
181 |
lemma "(\<forall>x\<in>A. P x) = (\<forall>x. x\<in>A \<longrightarrow> P x)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
182 |
by blast |
10294 | 183 |
|
67406 | 184 |
text\<open> |
10294 | 185 |
@{thm[display] ballI[no_vars]} |
186 |
\rulename{ballI} |
|
67406 | 187 |
\<close> |
10294 | 188 |
|
67406 | 189 |
text\<open> |
10294 | 190 |
@{thm[display] bspec[no_vars]} |
191 |
\rulename{bspec} |
|
67406 | 192 |
\<close> |
10294 | 193 |
|
67406 | 194 |
text\<open>indexed unions and variations\<close> |
10294 | 195 |
|
196 |
lemma "(\<Union>x. B x) = (\<Union>x\<in>UNIV. B x)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
197 |
by blast |
10294 | 198 |
|
67406 | 199 |
text\<open> |
10294 | 200 |
@{thm[display] UN_iff[no_vars]} |
201 |
\rulename{UN_iff} |
|
67406 | 202 |
\<close> |
10294 | 203 |
|
67406 | 204 |
text\<open> |
10294 | 205 |
@{thm[display] Union_iff[no_vars]} |
206 |
\rulename{Union_iff} |
|
67406 | 207 |
\<close> |
10294 | 208 |
|
209 |
lemma "(\<Union>x\<in>A. B x) = {y. \<exists>x\<in>A. y \<in> B x}" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
210 |
by blast |
10294 | 211 |
|
212 |
lemma "\<Union>S = (\<Union>x\<in>S. x)" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
213 |
by blast |
10294 | 214 |
|
67406 | 215 |
text\<open> |
10294 | 216 |
@{thm[display] UN_I[no_vars]} |
217 |
\rulename{UN_I} |
|
67406 | 218 |
\<close> |
10294 | 219 |
|
67406 | 220 |
text\<open> |
10294 | 221 |
@{thm[display] UN_E[no_vars]} |
222 |
\rulename{UN_E} |
|
67406 | 223 |
\<close> |
10294 | 224 |
|
67406 | 225 |
text\<open>indexed intersections\<close> |
10294 | 226 |
|
227 |
lemma "(\<Inter>x. B x) = {y. \<forall>x. y \<in> B x}" |
|
10864
f0b0a125ae4b
revisions corresponding to the new version of sets.tex
paulson
parents:
10341
diff
changeset
|
228 |
by blast |
10294 | 229 |
|
67406 | 230 |
text\<open> |
10294 | 231 |
@{thm[display] INT_iff[no_vars]} |
232 |
\rulename{INT_iff} |
|
67406 | 233 |
\<close> |
10294 | 234 |
|
67406 | 235 |
text\<open> |
10294 | 236 |
@{thm[display] Inter_iff[no_vars]} |
237 |
\rulename{Inter_iff} |
|
67406 | 238 |
\<close> |
10294 | 239 |
|
67406 | 240 |
text\<open>mention also card, Pow, etc.\<close> |
10294 | 241 |
|
242 |
||
67406 | 243 |
text\<open> |
10294 | 244 |
@{thm[display] card_Un_Int[no_vars]} |
245 |
\rulename{card_Un_Int} |
|
246 |
||
247 |
@{thm[display] card_Pow[no_vars]} |
|
248 |
\rulename{card_Pow} |
|
249 |
||
250 |
@{thm[display] n_subsets[no_vars]} |
|
251 |
\rulename{n_subsets} |
|
67406 | 252 |
\<close> |
10294 | 253 |
|
254 |
end |