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