| author | wenzelm | 
| Mon, 22 Mar 2010 19:23:03 +0100 | |
| changeset 35894 | ab6dc4d86ea1 | 
| parent 35416 | d8d7d1b785af | 
| child 38514 | bd9c4e8281ec | 
| permissions | -rw-r--r-- | 
| 1478 | 1 | (* Title: ZF/sum.thy | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 0 | 3 | Copyright 1993 University of Cambridge | 
| 4 | *) | |
| 5 | ||
| 13356 | 6 | header{*Disjoint Sums*}
 | 
| 7 | ||
| 16417 | 8 | theory Sum imports Bool equalities begin | 
| 3923 | 9 | |
| 13356 | 10 | text{*And the "Part" primitive for simultaneous recursive type definitions*}
 | 
| 11 | ||
| 3923 | 12 | global | 
| 13 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 14 | definition sum :: "[i,i]=>i" (infixr "+" 65) where | 
| 13240 | 15 |      "A+B == {0}*A Un {1}*B"
 | 
| 16 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 17 | definition Inl :: "i=>i" where | 
| 13240 | 18 | "Inl(a) == <0,a>" | 
| 19 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 20 | definition Inr :: "i=>i" where | 
| 13240 | 21 | "Inr(b) == <1,b>" | 
| 22 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 23 | definition "case" :: "[i=>i, i=>i, i]=>i" where | 
| 13240 | 24 | "case(c,d) == (%<y,z>. cond(y, d(z), c(z)))" | 
| 25 | ||
| 26 | (*operator for selecting out the various summands*) | |
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 27 | definition Part :: "[i,i=>i] => i" where | 
| 13240 | 28 |      "Part(A,h) == {x: A. EX z. x = h(z)}"
 | 
| 0 | 29 | |
| 3940 | 30 | local | 
| 3923 | 31 | |
| 13356 | 32 | subsection{*Rules for the @{term Part} Primitive*}
 | 
| 13240 | 33 | |
| 34 | lemma Part_iff: | |
| 35 | "a : Part(A,h) <-> a:A & (EX y. a=h(y))" | |
| 36 | apply (unfold Part_def) | |
| 37 | apply (rule separation) | |
| 38 | done | |
| 39 | ||
| 40 | lemma Part_eqI [intro]: | |
| 41 | "[| a : A; a=h(b) |] ==> a : Part(A,h)" | |
| 13255 | 42 | by (unfold Part_def, blast) | 
| 13240 | 43 | |
| 44 | lemmas PartI = refl [THEN [2] Part_eqI] | |
| 45 | ||
| 46 | lemma PartE [elim!]: | |
| 47 | "[| a : Part(A,h); !!z. [| a : A; a=h(z) |] ==> P | |
| 48 | |] ==> P" | |
| 13255 | 49 | apply (unfold Part_def, blast) | 
| 13240 | 50 | done | 
| 51 | ||
| 52 | lemma Part_subset: "Part(A,h) <= A" | |
| 53 | apply (unfold Part_def) | |
| 54 | apply (rule Collect_subset) | |
| 55 | done | |
| 56 | ||
| 57 | ||
| 13356 | 58 | subsection{*Rules for Disjoint Sums*}
 | 
| 13240 | 59 | |
| 60 | lemmas sum_defs = sum_def Inl_def Inr_def case_def | |
| 61 | ||
| 62 | lemma Sigma_bool: "Sigma(bool,C) = C(0) + C(1)" | |
| 13255 | 63 | by (unfold bool_def sum_def, blast) | 
| 13240 | 64 | |
| 65 | (** Introduction rules for the injections **) | |
| 66 | ||
| 67 | lemma InlI [intro!,simp,TC]: "a : A ==> Inl(a) : A+B" | |
| 13255 | 68 | by (unfold sum_defs, blast) | 
| 13240 | 69 | |
| 70 | lemma InrI [intro!,simp,TC]: "b : B ==> Inr(b) : A+B" | |
| 13255 | 71 | by (unfold sum_defs, blast) | 
| 13240 | 72 | |
| 73 | (** Elimination rules **) | |
| 74 | ||
| 75 | lemma sumE [elim!]: | |
| 76 | "[| u: A+B; | |
| 77 | !!x. [| x:A; u=Inl(x) |] ==> P; | |
| 78 | !!y. [| y:B; u=Inr(y) |] ==> P | |
| 79 | |] ==> P" | |
| 13255 | 80 | by (unfold sum_defs, blast) | 
| 13240 | 81 | |
| 82 | (** Injection and freeness equivalences, for rewriting **) | |
| 83 | ||
| 84 | lemma Inl_iff [iff]: "Inl(a)=Inl(b) <-> a=b" | |
| 13255 | 85 | by (simp add: sum_defs) | 
| 13240 | 86 | |
| 87 | lemma Inr_iff [iff]: "Inr(a)=Inr(b) <-> a=b" | |
| 13255 | 88 | by (simp add: sum_defs) | 
| 13240 | 89 | |
| 13823 | 90 | lemma Inl_Inr_iff [simp]: "Inl(a)=Inr(b) <-> False" | 
| 13255 | 91 | by (simp add: sum_defs) | 
| 13240 | 92 | |
| 13823 | 93 | lemma Inr_Inl_iff [simp]: "Inr(b)=Inl(a) <-> False" | 
| 13255 | 94 | by (simp add: sum_defs) | 
| 13240 | 95 | |
| 96 | lemma sum_empty [simp]: "0+0 = 0" | |
| 13255 | 97 | by (simp add: sum_defs) | 
| 13240 | 98 | |
| 99 | (*Injection and freeness rules*) | |
| 100 | ||
| 101 | lemmas Inl_inject = Inl_iff [THEN iffD1, standard] | |
| 102 | lemmas Inr_inject = Inr_iff [THEN iffD1, standard] | |
| 13823 | 103 | lemmas Inl_neq_Inr = Inl_Inr_iff [THEN iffD1, THEN FalseE, elim!] | 
| 104 | lemmas Inr_neq_Inl = Inr_Inl_iff [THEN iffD1, THEN FalseE, elim!] | |
| 13240 | 105 | |
| 106 | ||
| 107 | lemma InlD: "Inl(a): A+B ==> a: A" | |
| 13255 | 108 | by blast | 
| 13240 | 109 | |
| 110 | lemma InrD: "Inr(b): A+B ==> b: B" | |
| 13255 | 111 | by blast | 
| 13240 | 112 | |
| 113 | lemma sum_iff: "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))" | |
| 13255 | 114 | by blast | 
| 115 | ||
| 116 | lemma Inl_in_sum_iff [simp]: "(Inl(x) \<in> A+B) <-> (x \<in> A)"; | |
| 117 | by auto | |
| 118 | ||
| 119 | lemma Inr_in_sum_iff [simp]: "(Inr(y) \<in> A+B) <-> (y \<in> B)"; | |
| 120 | by auto | |
| 13240 | 121 | |
| 122 | lemma sum_subset_iff: "A+B <= C+D <-> A<=C & B<=D" | |
| 13255 | 123 | by blast | 
| 13240 | 124 | |
| 125 | lemma sum_equal_iff: "A+B = C+D <-> A=C & B=D" | |
| 13255 | 126 | by (simp add: extension sum_subset_iff, blast) | 
| 13240 | 127 | |
| 128 | lemma sum_eq_2_times: "A+A = 2*A" | |
| 13255 | 129 | by (simp add: sum_def, blast) | 
| 13240 | 130 | |
| 131 | ||
| 13356 | 132 | subsection{*The Eliminator: @{term case}*}
 | 
| 0 | 133 | |
| 13240 | 134 | lemma case_Inl [simp]: "case(c, d, Inl(a)) = c(a)" | 
| 13255 | 135 | by (simp add: sum_defs) | 
| 13240 | 136 | |
| 137 | lemma case_Inr [simp]: "case(c, d, Inr(b)) = d(b)" | |
| 13255 | 138 | by (simp add: sum_defs) | 
| 13240 | 139 | |
| 140 | lemma case_type [TC]: | |
| 141 | "[| u: A+B; | |
| 142 | !!x. x: A ==> c(x): C(Inl(x)); | |
| 143 | !!y. y: B ==> d(y): C(Inr(y)) | |
| 144 | |] ==> case(c,d,u) : C(u)" | |
| 13255 | 145 | by auto | 
| 13240 | 146 | |
| 147 | lemma expand_case: "u: A+B ==> | |
| 148 | R(case(c,d,u)) <-> | |
| 149 | ((ALL x:A. u = Inl(x) --> R(c(x))) & | |
| 150 | (ALL y:B. u = Inr(y) --> R(d(y))))" | |
| 151 | by auto | |
| 152 | ||
| 153 | lemma case_cong: | |
| 154 | "[| z: A+B; | |
| 155 | !!x. x:A ==> c(x)=c'(x); | |
| 156 | !!y. y:B ==> d(y)=d'(y) | |
| 157 | |] ==> case(c,d,z) = case(c',d',z)" | |
| 13255 | 158 | by auto | 
| 13240 | 159 | |
| 160 | lemma case_case: "z: A+B ==> | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
24893diff
changeset | 161 | case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) = | 
| 13240 | 162 | case(%x. c(c'(x)), %y. d(d'(y)), z)" | 
| 163 | by auto | |
| 164 | ||
| 165 | ||
| 13356 | 166 | subsection{*More Rules for @{term "Part(A,h)"}*}
 | 
| 13240 | 167 | |
| 168 | lemma Part_mono: "A<=B ==> Part(A,h)<=Part(B,h)" | |
| 13255 | 169 | by blast | 
| 13240 | 170 | |
| 171 | lemma Part_Collect: "Part(Collect(A,P), h) = Collect(Part(A,h), P)" | |
| 13255 | 172 | by blast | 
| 13240 | 173 | |
| 174 | lemmas Part_CollectE = | |
| 175 | Part_Collect [THEN equalityD1, THEN subsetD, THEN CollectE, standard] | |
| 176 | ||
| 177 | lemma Part_Inl: "Part(A+B,Inl) = {Inl(x). x: A}"
 | |
| 13255 | 178 | by blast | 
| 13240 | 179 | |
| 180 | lemma Part_Inr: "Part(A+B,Inr) = {Inr(y). y: B}"
 | |
| 13255 | 181 | by blast | 
| 13240 | 182 | |
| 183 | lemma PartD1: "a : Part(A,h) ==> a : A" | |
| 13255 | 184 | by (simp add: Part_def) | 
| 13240 | 185 | |
| 186 | lemma Part_id: "Part(A,%x. x) = A" | |
| 13255 | 187 | by blast | 
| 13240 | 188 | |
| 189 | lemma Part_Inr2: "Part(A+B, %x. Inr(h(x))) = {Inr(y). y: Part(B,h)}"
 | |
| 13255 | 190 | by blast | 
| 13240 | 191 | |
| 192 | lemma Part_sum_equality: "C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C" | |
| 13255 | 193 | by blast | 
| 13240 | 194 | |
| 0 | 195 | end |