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