doc-src/TutorialI/Sets/Examples.thy
author paulson
Thu, 26 Oct 2000 11:27:48 +0200
changeset 10341 6eb91805a012
parent 10294 2ec9c808a8a7
child 10864 f0b0a125ae4b
permissions -rw-r--r--
added the $Id:$ line
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10341
6eb91805a012 added the $Id:$ line
paulson
parents: 10294
diff changeset
     1
(* ID:         $Id$ *)
10294
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     2
theory Examples = Main:
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     3
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     4
ML "reset eta_contract"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     5
ML "Pretty.setmargin 64"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     6
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     7
text{*membership, intersection *}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     8
text{*difference and empty set*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
     9
text{*complement, union and universal set*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    10
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    11
lemma "(x \<in> A \<inter> B) = (x \<in> A \<and> x \<in> B)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    12
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    13
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    14
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    15
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    16
@{thm[display] IntI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    17
\rulename{IntI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    18
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    19
@{thm[display] IntD1[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    20
\rulename{IntD1}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    21
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    22
@{thm[display] IntD2[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    23
\rulename{IntD2}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    24
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    25
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    26
lemma "(x \<in> -A) = (x \<notin> A)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    27
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    28
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    29
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    30
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    31
@{thm[display] Compl_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    32
\rulename{Compl_iff}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    33
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    34
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    35
lemma "- (A \<union> B) = -A \<inter> -B"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    36
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    37
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    38
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    39
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    40
@{thm[display] Compl_Un[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    41
\rulename{Compl_Un}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    42
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    43
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    44
lemma "A-A = {}"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    45
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    46
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    47
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    48
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    49
@{thm[display] Diff_disjoint[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    50
\rulename{Diff_disjoint}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    51
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    52
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    53
  
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    54
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    55
lemma "A \<union> -A = UNIV"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    56
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    57
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    58
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    59
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    60
@{thm[display] Compl_partition[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    61
\rulename{Compl_partition}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    62
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    63
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    64
text{*subset relation*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    65
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    66
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    67
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    68
@{thm[display] subsetI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    69
\rulename{subsetI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    70
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    71
@{thm[display] subsetD[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    72
\rulename{subsetD}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    73
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    74
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    75
lemma "((A \<union> B) \<subseteq> C) = (A \<subseteq> C \<and> B \<subseteq> C)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    76
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    77
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    78
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    79
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    80
@{thm[display] Un_subset_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    81
\rulename{Un_subset_iff}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    82
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    83
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    84
lemma "(A \<subseteq> -B) = (B \<subseteq> -A)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    85
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    86
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    87
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    88
lemma "(A <= -B) = (B <= -A)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    89
  oops
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    90
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    91
text{*ASCII version: blast fails because of overloading because
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    92
 it doesn't have to be sets*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    93
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    94
lemma "((A:: 'a set) <= -B) = (B <= -A)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    95
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    96
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    97
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    98
text{*A type constraint lets it work*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
    99
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   100
text{*An issue here: how do we discuss the distinction between ASCII and
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   101
X-symbol notation?  Here the latter disambiguates.*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   102
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   103
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   104
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   105
set extensionality
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   106
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   107
@{thm[display] set_ext[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   108
\rulename{set_ext}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   109
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   110
@{thm[display] equalityI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   111
\rulename{equalityI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   112
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   113
@{thm[display] equalityE[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   114
\rulename{equalityE}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   115
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   116
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   117
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   118
text{*finite sets: insertion and membership relation*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   119
text{*finite set notation*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   120
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   121
lemma "insert x A = {x} \<union> A"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   122
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   123
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   124
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   125
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   126
@{thm[display] insert_is_Un[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   127
\rulename{insert_is_Un}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   128
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   129
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   130
lemma "{a,b} \<union> {c,d} = {a,b,c,d}"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   131
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   132
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   133
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   134
lemma "{a,b} \<inter> {b,c} = {b}"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   135
  apply (auto)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   136
  oops
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   137
text{*fails because it isn't valid*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   138
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   139
lemma "{a,b} \<inter> {b,c} = (if a=c then {a,b} else {b})"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   140
  apply (simp)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   141
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   142
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   143
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   144
text{*or just force or auto.  blast alone can't handle the if-then-else*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   145
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   146
text{*next: some comprehension examples*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   147
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   148
lemma "(a \<in> {z. P z}) = P a"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   149
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   150
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   151
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   152
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   153
@{thm[display] mem_Collect_eq[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   154
\rulename{mem_Collect_eq}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   155
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   156
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   157
lemma "{x. x \<in> A} = A"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   158
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   159
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   160
  
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   161
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   162
@{thm[display] Collect_mem_eq[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   163
\rulename{Collect_mem_eq}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   164
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   165
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   166
lemma "{x. P x \<or> x \<in> A} = {x. P x} \<union> A"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   167
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   168
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   169
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   170
lemma "{x. P x \<longrightarrow> Q x} = -{x. P x} \<union> {x. Q x}"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   171
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   172
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   173
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   174
constdefs
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   175
  prime   :: "nat set"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   176
    "prime == {p. 1<p & (ALL m. m dvd p --> m=1 | m=p)}"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   177
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   178
lemma "{p*q | p q. p\<in>prime \<and> q\<in>prime} = 
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   179
       {z. \<exists>p q. z = p*q \<and> p\<in>prime \<and> q\<in>prime}"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   180
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   181
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   182
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   183
text{*binders*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   184
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   185
text{*bounded quantifiers*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   186
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   187
lemma "(\<exists>x\<in>A. P x) = (\<exists>x. x\<in>A \<and> P x)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   188
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   189
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   190
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   191
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   192
@{thm[display] bexI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   193
\rulename{bexI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   194
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   195
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   196
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   197
@{thm[display] bexE[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   198
\rulename{bexE}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   199
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   200
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   201
lemma "(\<forall>x\<in>A. P x) = (\<forall>x. x\<in>A \<longrightarrow> P x)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   202
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   203
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   204
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   205
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   206
@{thm[display] ballI[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   207
\rulename{ballI}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   208
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   209
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   210
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   211
@{thm[display] bspec[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   212
\rulename{bspec}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   213
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   214
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   215
text{*indexed unions and variations*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   216
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   217
lemma "(\<Union>x. B x) = (\<Union>x\<in>UNIV. B x)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   218
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   219
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   220
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   221
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   222
@{thm[display] UN_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   223
\rulename{UN_iff}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   224
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   225
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   226
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   227
@{thm[display] Union_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   228
\rulename{Union_iff}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   229
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   230
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   231
lemma "(\<Union>x\<in>A. B x) = {y. \<exists>x\<in>A. y \<in> B x}"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   232
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   233
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   234
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   235
lemma "\<Union>S = (\<Union>x\<in>S. x)"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   236
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   237
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   238
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   239
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   240
@{thm[display] UN_I[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   241
\rulename{UN_I}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   242
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   243
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   244
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   245
@{thm[display] UN_E[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   246
\rulename{UN_E}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   247
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   248
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   249
text{*indexed intersections*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   250
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   251
lemma "(\<Inter>x. B x) = {y. \<forall>x. y \<in> B x}"
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   252
  apply (blast)
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   253
  done
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   254
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   255
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   256
@{thm[display] INT_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   257
\rulename{INT_iff}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   258
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   259
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   260
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   261
@{thm[display] Inter_iff[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   262
\rulename{Inter_iff}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   263
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   264
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   265
text{*mention also card, Pow, etc.*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   266
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   267
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   268
text{*
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   269
@{thm[display] card_Un_Int[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   270
\rulename{card_Un_Int}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   271
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   272
@{thm[display] card_Pow[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   273
\rulename{card_Pow}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   274
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   275
@{thm[display] n_subsets[no_vars]}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   276
\rulename{n_subsets}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   277
*}
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   278
2ec9c808a8a7 the Sets chapter and theories
paulson
parents:
diff changeset
   279
end