src/Doc/Tutorial/Sets/Examples.thy
author wenzelm
Wed, 12 Mar 2025 11:39:00 +0100
changeset 82265 4b875a4c83b0
parent 67613 ce654b0e6d69
permissions -rw-r--r--
update for release;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     2
39128
93a7365fb4ee turned eta_contract into proper configuration option;
wenzelm
parents: 38798
diff changeset
     3
declare [[eta_contract = false]]
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     4
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
     5
text\<open>membership, intersection\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
     6
text\<open>difference and empty set\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
     7
text\<open>complement, union and universal set\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     8
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    11
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    12
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    13
@{thm[display] IntI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    14
\rulename{IntI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    15
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    16
@{thm[display] IntD1[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    17
\rulename{IntD1}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    18
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    19
@{thm[display] IntD2[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    20
\rulename{IntD2}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    21
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    22
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    25
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    26
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    27
@{thm[display] Compl_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    28
\rulename{Compl_iff}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    29
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    30
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    33
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    34
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    35
@{thm[display] Compl_Un[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    36
\rulename{Compl_Un}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    37
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    38
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    39
lemma "A-A = {}"
10864
f0b0a125ae4b revisions corresponding to the new version of sets.tex
paulson
parents: 10341
diff changeset
    40
by blast
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    41
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    42
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    43
@{thm[display] Diff_disjoint[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    44
\rulename{Diff_disjoint}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    45
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    46
55159
608c157d743d Replacing the theory Library/Binomial by Number_Theory/Binomial
paulson <lp15@cam.ac.uk>
parents: 48985
diff changeset
    47
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    48
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    51
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    52
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    53
@{thm[display] Compl_partition[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    54
\rulename{Compl_partition}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    55
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    56
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    57
text\<open>subset relation\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    58
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    59
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    60
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    61
@{thm[display] subsetI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    62
\rulename{subsetI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    63
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    64
@{thm[display] subsetD[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    65
\rulename{subsetD}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    66
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    67
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    70
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    71
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    72
@{thm[display] Un_subset_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    73
\rulename{Un_subset_iff}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    74
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    75
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    78
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    79
lemma "(A <= -B) = (B <= -A)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    80
  oops
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    81
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    82
text\<open>ASCII version: blast fails because of overloading because
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    83
 it doesn't have to be sets\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    84
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    87
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    88
text\<open>A type constraint lets it work\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    89
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    90
text\<open>An issue here: how do we discuss the distinction between ASCII and
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    91
symbol notation?  Here the latter disambiguates.\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    92
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    93
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
    94
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    95
set extensionality
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    96
39795
9e59b4c11039 updated files to recent changes
haftmann
parents: 39128
diff changeset
    97
@{thm[display] set_eqI[no_vars]}
9e59b4c11039 updated files to recent changes
haftmann
parents: 39128
diff changeset
    98
\rulename{set_eqI}
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    99
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   100
@{thm[display] equalityI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   101
\rulename{equalityI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   102
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   103
@{thm[display] equalityE[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   104
\rulename{equalityE}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   105
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   106
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   107
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   108
text\<open>finite sets: insertion and membership relation\<close>
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   109
text\<open>finite set notation\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   110
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   113
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   114
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   115
@{thm[display] insert_is_Un[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   116
\rulename{insert_is_Un}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   117
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   118
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   121
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   125
text\<open>fails because it isn't valid\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   126
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   130
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   131
text\<open>or just force or auto.  blast alone can't handle the if-then-else\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   132
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   133
text\<open>next: some comprehension examples\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   134
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   137
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   138
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   139
@{thm[display] mem_Collect_eq[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   140
\rulename{mem_Collect_eq}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   141
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   142
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   146
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   147
@{thm[display] Collect_mem_eq[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   148
\rulename{Collect_mem_eq}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   149
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   150
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   153
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
   158
    "prime == {p. 1<p & (\<forall>m. m dvd p \<longrightarrow> m=1 \<or> m=p)}"
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   163
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   164
text\<open>binders\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   165
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   166
text\<open>bounded quantifiers\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   167
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   170
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   171
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   172
@{thm[display] bexI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   173
\rulename{bexI}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   174
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   175
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   176
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   177
@{thm[display] bexE[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   178
\rulename{bexE}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   179
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   180
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   183
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   184
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   185
@{thm[display] ballI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   186
\rulename{ballI}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   187
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   188
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   189
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   190
@{thm[display] bspec[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   191
\rulename{bspec}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   192
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   193
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   194
text\<open>indexed unions and variations\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   195
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   198
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   199
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   200
@{thm[display] UN_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   201
\rulename{UN_iff}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   202
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   203
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   204
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   205
@{thm[display] Union_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   206
\rulename{Union_iff}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   207
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   208
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   211
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   214
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   215
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   216
@{thm[display] UN_I[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   217
\rulename{UN_I}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   218
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   219
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   220
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   221
@{thm[display] UN_E[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   222
\rulename{UN_E}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   223
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   224
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   225
text\<open>indexed intersections\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   226
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   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
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   229
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   230
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   231
@{thm[display] INT_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   232
\rulename{INT_iff}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   233
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   234
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   235
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   236
@{thm[display] Inter_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   237
\rulename{Inter_iff}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   238
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   239
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   240
text\<open>mention also card, Pow, etc.\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   241
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   242
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   243
text\<open>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   244
@{thm[display] card_Un_Int[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   245
\rulename{card_Un_Int}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   246
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   247
@{thm[display] card_Pow[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   248
\rulename{card_Pow}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   249
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   250
@{thm[display] n_subsets[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   251
\rulename{n_subsets}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 59667
diff changeset
   252
\<close>
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   253
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   254
end