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