| author | blanchet | 
| Thu, 07 Aug 2014 12:17:41 +0200 | |
| changeset 57816 | d8bbb97689d3 | 
| parent 55931 | 62156e694f3d | 
| child 58189 | 9d714be4f028 | 
| permissions | -rw-r--r-- | 
| 10213 | 1 | (* Title: HOL/Sum_Type.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1992 University of Cambridge | |
| 4 | *) | |
| 5 | ||
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 6 | header{*The Disjoint Sum of Two Types*}
 | 
| 10213 | 7 | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 8 | theory Sum_Type | 
| 33961 | 9 | imports Typedef Inductive Fun | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 10 | begin | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 11 | |
| 33962 | 12 | subsection {* Construction of the sum type and its basic abstract operations *}
 | 
| 10213 | 13 | |
| 33962 | 14 | definition Inl_Rep :: "'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool \<Rightarrow> bool" where | 
| 15 | "Inl_Rep a x y p \<longleftrightarrow> x = a \<and> p" | |
| 10213 | 16 | |
| 33962 | 17 | definition Inr_Rep :: "'b \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> bool \<Rightarrow> bool" where | 
| 18 | "Inr_Rep b x y p \<longleftrightarrow> y = b \<and> \<not> p" | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 19 | |
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45204diff
changeset | 20 | definition "sum = {f. (\<exists>a. f = Inl_Rep (a::'a)) \<or> (\<exists>b. f = Inr_Rep (b::'b))}"
 | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45204diff
changeset | 21 | |
| 49834 | 22 | typedef ('a, 'b) sum (infixr "+" 10) = "sum :: ('a => 'b => bool => bool) set"
 | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45204diff
changeset | 23 | unfolding sum_def by auto | 
| 10213 | 24 | |
| 37388 | 25 | lemma Inl_RepI: "Inl_Rep a \<in> sum" | 
| 26 | by (auto simp add: sum_def) | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 27 | |
| 37388 | 28 | lemma Inr_RepI: "Inr_Rep b \<in> sum" | 
| 29 | by (auto simp add: sum_def) | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 30 | |
| 37388 | 31 | lemma inj_on_Abs_sum: "A \<subseteq> sum \<Longrightarrow> inj_on Abs_sum A" | 
| 32 | by (rule inj_on_inverseI, rule Abs_sum_inverse) auto | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 33 | |
| 33962 | 34 | lemma Inl_Rep_inject: "inj_on Inl_Rep A" | 
| 35 | proof (rule inj_onI) | |
| 36 | show "\<And>a c. Inl_Rep a = Inl_Rep c \<Longrightarrow> a = c" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 37 | by (auto simp add: Inl_Rep_def fun_eq_iff) | 
| 33962 | 38 | qed | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 39 | |
| 33962 | 40 | lemma Inr_Rep_inject: "inj_on Inr_Rep A" | 
| 41 | proof (rule inj_onI) | |
| 42 | show "\<And>b d. Inr_Rep b = Inr_Rep d \<Longrightarrow> b = d" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 43 | by (auto simp add: Inr_Rep_def fun_eq_iff) | 
| 33962 | 44 | qed | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 45 | |
| 33962 | 46 | lemma Inl_Rep_not_Inr_Rep: "Inl_Rep a \<noteq> Inr_Rep b" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 47 | by (auto simp add: Inl_Rep_def Inr_Rep_def fun_eq_iff) | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 48 | |
| 33962 | 49 | definition Inl :: "'a \<Rightarrow> 'a + 'b" where | 
| 37388 | 50 | "Inl = Abs_sum \<circ> Inl_Rep" | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 51 | |
| 33962 | 52 | definition Inr :: "'b \<Rightarrow> 'a + 'b" where | 
| 37388 | 53 | "Inr = Abs_sum \<circ> Inr_Rep" | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 54 | |
| 29025 
8c8859c0d734
move lemmas from Numeral_Type.thy to other theories
 huffman parents: 
28524diff
changeset | 55 | lemma inj_Inl [simp]: "inj_on Inl A" | 
| 37388 | 56 | by (auto simp add: Inl_def intro!: comp_inj_on Inl_Rep_inject inj_on_Abs_sum Inl_RepI) | 
| 29025 
8c8859c0d734
move lemmas from Numeral_Type.thy to other theories
 huffman parents: 
28524diff
changeset | 57 | |
| 33962 | 58 | lemma Inl_inject: "Inl x = Inl y \<Longrightarrow> x = y" | 
| 59 | using inj_Inl by (rule injD) | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 60 | |
| 29025 
8c8859c0d734
move lemmas from Numeral_Type.thy to other theories
 huffman parents: 
28524diff
changeset | 61 | lemma inj_Inr [simp]: "inj_on Inr A" | 
| 37388 | 62 | by (auto simp add: Inr_def intro!: comp_inj_on Inr_Rep_inject inj_on_Abs_sum Inr_RepI) | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 63 | |
| 33962 | 64 | lemma Inr_inject: "Inr x = Inr y \<Longrightarrow> x = y" | 
| 65 | using inj_Inr by (rule injD) | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 66 | |
| 33962 | 67 | lemma Inl_not_Inr: "Inl a \<noteq> Inr b" | 
| 68 | proof - | |
| 37388 | 69 |   from Inl_RepI [of a] Inr_RepI [of b] have "{Inl_Rep a, Inr_Rep b} \<subseteq> sum" by auto
 | 
| 70 |   with inj_on_Abs_sum have "inj_on Abs_sum {Inl_Rep a, Inr_Rep b}" .
 | |
| 71 | with Inl_Rep_not_Inr_Rep [of a b] inj_on_contraD have "Abs_sum (Inl_Rep a) \<noteq> Abs_sum (Inr_Rep b)" by auto | |
| 33962 | 72 | then show ?thesis by (simp add: Inl_def Inr_def) | 
| 73 | qed | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 74 | |
| 33962 | 75 | lemma Inr_not_Inl: "Inr b \<noteq> Inl a" | 
| 76 | using Inl_not_Inr by (rule not_sym) | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 77 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 78 | lemma sumE: | 
| 33962 | 79 | assumes "\<And>x::'a. s = Inl x \<Longrightarrow> P" | 
| 80 | and "\<And>y::'b. s = Inr y \<Longrightarrow> P" | |
| 81 | shows P | |
| 37388 | 82 | proof (rule Abs_sum_cases [of s]) | 
| 33962 | 83 | fix f | 
| 37388 | 84 | assume "s = Abs_sum f" and "f \<in> sum" | 
| 85 | with assms show P by (auto simp add: sum_def Inl_def Inr_def) | |
| 33962 | 86 | qed | 
| 33961 | 87 | |
| 55469 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 blanchet parents: 
55468diff
changeset | 88 | free_constructors case_sum for | 
| 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 blanchet parents: 
55468diff
changeset | 89 | isl: Inl projl | 
| 
b19dd108f0c2
aligned the syntax for 'free_constructors' on the 'datatype_new' and 'codatatype' syntax
 blanchet parents: 
55468diff
changeset | 90 | | Inr projr | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 91 | by (erule sumE, assumption) (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr) | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 92 | |
| 55442 | 93 | text {* Avoid name clashes by prefixing the output of @{text rep_datatype} with @{text old}. *}
 | 
| 94 | ||
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 95 | setup {* Sign.mandatory_path "old" *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 96 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
37388diff
changeset | 97 | rep_datatype Inl Inr | 
| 33961 | 98 | proof - | 
| 99 | fix P | |
| 100 | fix s :: "'a + 'b" | |
| 101 | assume x: "\<And>x\<Colon>'a. P (Inl x)" and y: "\<And>y\<Colon>'b. P (Inr y)" | |
| 102 | then show "P s" by (auto intro: sumE [of s]) | |
| 33962 | 103 | qed (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr) | 
| 104 | ||
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 105 | setup {* Sign.parent_path *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 106 | |
| 55468 
98b25c51e9e5
renamed 'wrap_free_constructors' to 'free_constructors' (cf. 'functor', 'bnf', etc.)
 blanchet parents: 
55467diff
changeset | 107 | text {* But erase the prefix for properties that are not generated by @{text free_constructors}. *}
 | 
| 55442 | 108 | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 109 | setup {* Sign.mandatory_path "sum" *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 110 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 111 | declare | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 112 | old.sum.inject[iff del] | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 113 | old.sum.distinct(1)[simp del, induct_simp del] | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 114 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 115 | lemmas induct = old.sum.induct | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 116 | lemmas inducts = old.sum.inducts | 
| 55642 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55575diff
changeset | 117 | lemmas rec = old.sum.rec | 
| 
63beb38e9258
adapted to renaming of datatype 'cases' and 'recs' to 'case' and 'rec'
 blanchet parents: 
55575diff
changeset | 118 | lemmas simps = sum.inject sum.distinct sum.case sum.rec | 
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 119 | |
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 120 | setup {* Sign.parent_path *}
 | 
| 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 121 | |
| 55931 | 122 | primrec map_sum :: "('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> 'a + 'b \<Rightarrow> 'c + 'd" where
 | 
| 123 | "map_sum f1 f2 (Inl a) = Inl (f1 a)" | |
| 124 | | "map_sum f1 f2 (Inr a) = Inr (f2 a)" | |
| 40610 | 125 | |
| 55931 | 126 | functor map_sum: map_sum proof - | 
| 41372 | 127 | fix f g h i | 
| 55931 | 128 | show "map_sum f g \<circ> map_sum h i = map_sum (f \<circ> h) (g \<circ> i)" | 
| 41372 | 129 | proof | 
| 130 | fix s | |
| 55931 | 131 | show "(map_sum f g \<circ> map_sum h i) s = map_sum (f \<circ> h) (g \<circ> i) s" | 
| 41372 | 132 | by (cases s) simp_all | 
| 133 | qed | |
| 40610 | 134 | next | 
| 135 | fix s | |
| 55931 | 136 | show "map_sum id id = id" | 
| 41372 | 137 | proof | 
| 138 | fix s | |
| 55931 | 139 | show "map_sum id id s = id s" | 
| 41372 | 140 | by (cases s) simp_all | 
| 141 | qed | |
| 40610 | 142 | qed | 
| 143 | ||
| 53010 | 144 | lemma split_sum_all: "(\<forall>x. P x) \<longleftrightarrow> (\<forall>x. P (Inl x)) \<and> (\<forall>x. P (Inr x))" | 
| 145 | by (auto intro: sum.induct) | |
| 146 | ||
| 147 | lemma split_sum_ex: "(\<exists>x. P x) \<longleftrightarrow> (\<exists>x. P (Inl x)) \<or> (\<exists>x. P (Inr x))" | |
| 148 | using split_sum_all[of "\<lambda>x. \<not>P x"] by blast | |
| 33961 | 149 | |
| 33962 | 150 | subsection {* Projections *}
 | 
| 151 | ||
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 152 | lemma case_sum_KK [simp]: "case_sum (\<lambda>x. a) (\<lambda>x. a) = (\<lambda>x. a)" | 
| 33961 | 153 | by (rule ext) (simp split: sum.split) | 
| 154 | ||
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 155 | lemma surjective_sum: "case_sum (\<lambda>x::'a. f (Inl x)) (\<lambda>y::'b. f (Inr y)) = f" | 
| 33962 | 156 | proof | 
| 157 | fix s :: "'a + 'b" | |
| 158 | show "(case s of Inl (x\<Colon>'a) \<Rightarrow> f (Inl x) | Inr (y\<Colon>'b) \<Rightarrow> f (Inr y)) = f s" | |
| 159 | by (cases s) simp_all | |
| 160 | qed | |
| 33961 | 161 | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 162 | lemma case_sum_inject: | 
| 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 163 | assumes a: "case_sum f1 f2 = case_sum g1 g2" | 
| 33962 | 164 | assumes r: "f1 = g1 \<Longrightarrow> f2 = g2 \<Longrightarrow> P" | 
| 165 | shows P | |
| 166 | proof (rule r) | |
| 167 | show "f1 = g1" proof | |
| 168 | fix x :: 'a | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 169 | from a have "case_sum f1 f2 (Inl x) = case_sum g1 g2 (Inl x)" by simp | 
| 33962 | 170 | then show "f1 x = g1 x" by simp | 
| 171 | qed | |
| 172 | show "f2 = g2" proof | |
| 173 | fix y :: 'b | |
| 55414 
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
 blanchet parents: 
55403diff
changeset | 174 | from a have "case_sum f1 f2 (Inr y) = case_sum g1 g2 (Inr y)" by simp | 
| 33962 | 175 | then show "f2 y = g2 y" by simp | 
| 176 | qed | |
| 177 | qed | |
| 178 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55534diff
changeset | 179 | primrec Suml :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a + 'b \<Rightarrow> 'c" where
 | 
| 33962 | 180 | "Suml f (Inl x) = f x" | 
| 181 | ||
| 55575 
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
 blanchet parents: 
55534diff
changeset | 182 | primrec Sumr :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a + 'b \<Rightarrow> 'c" where
 | 
| 33962 | 183 | "Sumr f (Inr x) = f x" | 
| 184 | ||
| 185 | lemma Suml_inject: | |
| 186 | assumes "Suml f = Suml g" shows "f = g" | |
| 187 | proof | |
| 188 | fix x :: 'a | |
| 189 | let ?s = "Inl x \<Colon> 'a + 'b" | |
| 190 | from assms have "Suml f ?s = Suml g ?s" by simp | |
| 191 | then show "f x = g x" by simp | |
| 33961 | 192 | qed | 
| 193 | ||
| 33962 | 194 | lemma Sumr_inject: | 
| 195 | assumes "Sumr f = Sumr g" shows "f = g" | |
| 196 | proof | |
| 197 | fix x :: 'b | |
| 198 | let ?s = "Inr x \<Colon> 'a + 'b" | |
| 199 | from assms have "Sumr f ?s = Sumr g ?s" by simp | |
| 200 | then show "f x = g x" by simp | |
| 201 | qed | |
| 33961 | 202 | |
| 33995 | 203 | |
| 33962 | 204 | subsection {* The Disjoint Sum of Sets *}
 | 
| 33961 | 205 | |
| 33962 | 206 | definition Plus :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a + 'b) set" (infixr "<+>" 65) where
 | 
| 207 | "A <+> B = Inl ` A \<union> Inr ` B" | |
| 208 | ||
| 40271 | 209 | hide_const (open) Plus --"Valuable identifier" | 
| 210 | ||
| 33962 | 211 | lemma InlI [intro!]: "a \<in> A \<Longrightarrow> Inl a \<in> A <+> B" | 
| 212 | by (simp add: Plus_def) | |
| 33961 | 213 | |
| 33962 | 214 | lemma InrI [intro!]: "b \<in> B \<Longrightarrow> Inr b \<in> A <+> B" | 
| 215 | by (simp add: Plus_def) | |
| 33961 | 216 | |
| 33962 | 217 | text {* Exhaustion rule for sums, a degenerate form of induction *}
 | 
| 218 | ||
| 219 | lemma PlusE [elim!]: | |
| 220 | "u \<in> A <+> B \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> u = Inl x \<Longrightarrow> P) \<Longrightarrow> (\<And>y. y \<in> B \<Longrightarrow> u = Inr y \<Longrightarrow> P) \<Longrightarrow> P" | |
| 221 | by (auto simp add: Plus_def) | |
| 33961 | 222 | |
| 33962 | 223 | lemma Plus_eq_empty_conv [simp]: "A <+> B = {} \<longleftrightarrow> A = {} \<and> B = {}"
 | 
| 224 | by auto | |
| 33961 | 225 | |
| 33962 | 226 | lemma UNIV_Plus_UNIV [simp]: "UNIV <+> UNIV = UNIV" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39198diff
changeset | 227 | proof (rule set_eqI) | 
| 33962 | 228 | fix u :: "'a + 'b" | 
| 229 | show "u \<in> UNIV <+> UNIV \<longleftrightarrow> u \<in> UNIV" by (cases u) auto | |
| 230 | qed | |
| 33961 | 231 | |
| 49948 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 232 | lemma UNIV_sum: | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 233 | "UNIV = Inl ` UNIV \<union> Inr ` UNIV" | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 234 | proof - | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 235 |   { fix x :: "'a + 'b"
 | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 236 | assume "x \<notin> range Inr" | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 237 | then have "x \<in> range Inl" | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 238 | by (cases x) simp_all | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 239 | } then show ?thesis by auto | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 240 | qed | 
| 
744934b818c7
moved quite generic material from theory Enum to more appropriate places
 haftmann parents: 
49834diff
changeset | 241 | |
| 55393 
ce5cebfaedda
 se 'wrap_free_constructors' to register 'sum' , 'prod', 'unit', 'bool' with their discriminators/selectors
 blanchet parents: 
53010diff
changeset | 242 | hide_const (open) Suml Sumr sum | 
| 45204 
5e4a1270c000
hide typedef-generated constants Product_Type.prod and Sum_Type.sum
 huffman parents: 
41505diff
changeset | 243 | |
| 10213 | 244 | end |