| author | bulwahn | 
| Mon, 23 Aug 2010 16:47:55 +0200 | |
| changeset 38664 | 7215ae18f44b | 
| parent 37678 | 0040bafffdef | 
| child 39198 | f967a16dfcdd | 
| 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: 
11609 
diff
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: 
11609 
diff
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: 
11609 
diff
changeset
 | 
10  | 
begin  | 
| 
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
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: 
11609 
diff
changeset
 | 
19  | 
|
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
37388 
diff
changeset
 | 
20  | 
typedef ('a, 'b) sum (infixr "+" 10) = "{f. (\<exists>a. f = Inl_Rep (a::'a)) \<or> (\<exists>b. f = Inr_Rep (b::'b))}"
 | 
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
21  | 
by auto  | 
| 10213 | 22  | 
|
| 37388 | 23  | 
lemma Inl_RepI: "Inl_Rep a \<in> sum"  | 
24  | 
by (auto simp add: sum_def)  | 
|
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
25  | 
|
| 37388 | 26  | 
lemma Inr_RepI: "Inr_Rep b \<in> sum"  | 
27  | 
by (auto simp add: sum_def)  | 
|
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
28  | 
|
| 37388 | 29  | 
lemma inj_on_Abs_sum: "A \<subseteq> sum \<Longrightarrow> inj_on Abs_sum A"  | 
30  | 
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: 
11609 
diff
changeset
 | 
31  | 
|
| 33962 | 32  | 
lemma Inl_Rep_inject: "inj_on Inl_Rep A"  | 
33  | 
proof (rule inj_onI)  | 
|
34  | 
show "\<And>a c. Inl_Rep a = Inl_Rep c \<Longrightarrow> a = c"  | 
|
35  | 
by (auto simp add: Inl_Rep_def expand_fun_eq)  | 
|
36  | 
qed  | 
|
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
37  | 
|
| 33962 | 38  | 
lemma Inr_Rep_inject: "inj_on Inr_Rep A"  | 
39  | 
proof (rule inj_onI)  | 
|
40  | 
show "\<And>b d. Inr_Rep b = Inr_Rep d \<Longrightarrow> b = d"  | 
|
41  | 
by (auto simp add: Inr_Rep_def expand_fun_eq)  | 
|
42  | 
qed  | 
|
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
43  | 
|
| 33962 | 44  | 
lemma Inl_Rep_not_Inr_Rep: "Inl_Rep a \<noteq> Inr_Rep b"  | 
45  | 
by (auto simp add: Inl_Rep_def Inr_Rep_def expand_fun_eq)  | 
|
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
46  | 
|
| 33962 | 47  | 
definition Inl :: "'a \<Rightarrow> 'a + 'b" where  | 
| 37388 | 48  | 
"Inl = Abs_sum \<circ> Inl_Rep"  | 
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
49  | 
|
| 33962 | 50  | 
definition Inr :: "'b \<Rightarrow> 'a + 'b" where  | 
| 37388 | 51  | 
"Inr = Abs_sum \<circ> Inr_Rep"  | 
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
52  | 
|
| 
29025
 
8c8859c0d734
move lemmas from Numeral_Type.thy to other theories
 
huffman 
parents: 
28524 
diff
changeset
 | 
53  | 
lemma inj_Inl [simp]: "inj_on Inl A"  | 
| 37388 | 54  | 
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: 
28524 
diff
changeset
 | 
55  | 
|
| 33962 | 56  | 
lemma Inl_inject: "Inl x = Inl y \<Longrightarrow> x = y"  | 
57  | 
using inj_Inl by (rule injD)  | 
|
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
58  | 
|
| 
29025
 
8c8859c0d734
move lemmas from Numeral_Type.thy to other theories
 
huffman 
parents: 
28524 
diff
changeset
 | 
59  | 
lemma inj_Inr [simp]: "inj_on Inr A"  | 
| 37388 | 60  | 
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: 
11609 
diff
changeset
 | 
61  | 
|
| 33962 | 62  | 
lemma Inr_inject: "Inr x = Inr y \<Longrightarrow> x = y"  | 
63  | 
using inj_Inr by (rule injD)  | 
|
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
64  | 
|
| 33962 | 65  | 
lemma Inl_not_Inr: "Inl a \<noteq> Inr b"  | 
66  | 
proof -  | 
|
| 37388 | 67  | 
  from Inl_RepI [of a] Inr_RepI [of b] have "{Inl_Rep a, Inr_Rep b} \<subseteq> sum" by auto
 | 
68  | 
  with inj_on_Abs_sum have "inj_on Abs_sum {Inl_Rep a, Inr_Rep b}" .
 | 
|
69  | 
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 | 70  | 
then show ?thesis by (simp add: Inl_def Inr_def)  | 
71  | 
qed  | 
|
| 
15391
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
72  | 
|
| 33962 | 73  | 
lemma Inr_not_Inl: "Inr b \<noteq> Inl a"  | 
74  | 
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: 
11609 
diff
changeset
 | 
75  | 
|
| 
 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 
paulson 
parents: 
11609 
diff
changeset
 | 
76  | 
lemma sumE:  | 
| 33962 | 77  | 
assumes "\<And>x::'a. s = Inl x \<Longrightarrow> P"  | 
78  | 
and "\<And>y::'b. s = Inr y \<Longrightarrow> P"  | 
|
79  | 
shows P  | 
|
| 37388 | 80  | 
proof (rule Abs_sum_cases [of s])  | 
| 33962 | 81  | 
fix f  | 
| 37388 | 82  | 
assume "s = Abs_sum f" and "f \<in> sum"  | 
83  | 
with assms show P by (auto simp add: sum_def Inl_def Inr_def)  | 
|
| 33962 | 84  | 
qed  | 
| 33961 | 85  | 
|
| 
37678
 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 
haftmann 
parents: 
37388 
diff
changeset
 | 
86  | 
rep_datatype Inl Inr  | 
| 33961 | 87  | 
proof -  | 
88  | 
fix P  | 
|
89  | 
fix s :: "'a + 'b"  | 
|
90  | 
assume x: "\<And>x\<Colon>'a. P (Inl x)" and y: "\<And>y\<Colon>'b. P (Inr y)"  | 
|
91  | 
then show "P s" by (auto intro: sumE [of s])  | 
|
| 33962 | 92  | 
qed (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr)  | 
93  | 
||
| 33961 | 94  | 
|
| 33962 | 95  | 
subsection {* Projections *}
 | 
96  | 
||
97  | 
lemma sum_case_KK [simp]: "sum_case (\<lambda>x. a) (\<lambda>x. a) = (\<lambda>x. a)"  | 
|
| 33961 | 98  | 
by (rule ext) (simp split: sum.split)  | 
99  | 
||
| 33962 | 100  | 
lemma surjective_sum: "sum_case (\<lambda>x::'a. f (Inl x)) (\<lambda>y::'b. f (Inr y)) = f"  | 
101  | 
proof  | 
|
102  | 
fix s :: "'a + 'b"  | 
|
103  | 
show "(case s of Inl (x\<Colon>'a) \<Rightarrow> f (Inl x) | Inr (y\<Colon>'b) \<Rightarrow> f (Inr y)) = f s"  | 
|
104  | 
by (cases s) simp_all  | 
|
105  | 
qed  | 
|
| 33961 | 106  | 
|
| 33962 | 107  | 
lemma sum_case_inject:  | 
108  | 
assumes a: "sum_case f1 f2 = sum_case g1 g2"  | 
|
109  | 
assumes r: "f1 = g1 \<Longrightarrow> f2 = g2 \<Longrightarrow> P"  | 
|
110  | 
shows P  | 
|
111  | 
proof (rule r)  | 
|
112  | 
show "f1 = g1" proof  | 
|
113  | 
fix x :: 'a  | 
|
114  | 
from a have "sum_case f1 f2 (Inl x) = sum_case g1 g2 (Inl x)" by simp  | 
|
115  | 
then show "f1 x = g1 x" by simp  | 
|
116  | 
qed  | 
|
117  | 
show "f2 = g2" proof  | 
|
118  | 
fix y :: 'b  | 
|
119  | 
from a have "sum_case f1 f2 (Inr y) = sum_case g1 g2 (Inr y)" by simp  | 
|
120  | 
then show "f2 y = g2 y" by simp  | 
|
121  | 
qed  | 
|
122  | 
qed  | 
|
123  | 
||
124  | 
lemma sum_case_weak_cong:  | 
|
125  | 
"s = t \<Longrightarrow> sum_case f g s = sum_case f g t"  | 
|
| 33961 | 126  | 
  -- {* Prevents simplification of @{text f} and @{text g}: much faster. *}
 | 
127  | 
by simp  | 
|
128  | 
||
| 33962 | 129  | 
primrec Projl :: "'a + 'b \<Rightarrow> 'a" where  | 
130  | 
Projl_Inl: "Projl (Inl x) = x"  | 
|
131  | 
||
132  | 
primrec Projr :: "'a + 'b \<Rightarrow> 'b" where  | 
|
133  | 
Projr_Inr: "Projr (Inr x) = x"  | 
|
134  | 
||
135  | 
primrec Suml :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a + 'b \<Rightarrow> 'c" where
 | 
|
136  | 
"Suml f (Inl x) = f x"  | 
|
137  | 
||
138  | 
primrec Sumr :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a + 'b \<Rightarrow> 'c" where
 | 
|
139  | 
"Sumr f (Inr x) = f x"  | 
|
140  | 
||
141  | 
lemma Suml_inject:  | 
|
142  | 
assumes "Suml f = Suml g" shows "f = g"  | 
|
143  | 
proof  | 
|
144  | 
fix x :: 'a  | 
|
145  | 
let ?s = "Inl x \<Colon> 'a + 'b"  | 
|
146  | 
from assms have "Suml f ?s = Suml g ?s" by simp  | 
|
147  | 
then show "f x = g x" by simp  | 
|
| 33961 | 148  | 
qed  | 
149  | 
||
| 33962 | 150  | 
lemma Sumr_inject:  | 
151  | 
assumes "Sumr f = Sumr g" shows "f = g"  | 
|
152  | 
proof  | 
|
153  | 
fix x :: 'b  | 
|
154  | 
let ?s = "Inr x \<Colon> 'a + 'b"  | 
|
155  | 
from assms have "Sumr f ?s = Sumr g ?s" by simp  | 
|
156  | 
then show "f x = g x" by simp  | 
|
157  | 
qed  | 
|
| 33961 | 158  | 
|
| 33995 | 159  | 
|
| 33962 | 160  | 
subsection {* The Disjoint Sum of Sets *}
 | 
| 33961 | 161  | 
|
| 33962 | 162  | 
definition Plus :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a + 'b) set" (infixr "<+>" 65) where
 | 
163  | 
"A <+> B = Inl ` A \<union> Inr ` B"  | 
|
164  | 
||
165  | 
lemma InlI [intro!]: "a \<in> A \<Longrightarrow> Inl a \<in> A <+> B"  | 
|
166  | 
by (simp add: Plus_def)  | 
|
| 33961 | 167  | 
|
| 33962 | 168  | 
lemma InrI [intro!]: "b \<in> B \<Longrightarrow> Inr b \<in> A <+> B"  | 
169  | 
by (simp add: Plus_def)  | 
|
| 33961 | 170  | 
|
| 33962 | 171  | 
text {* Exhaustion rule for sums, a degenerate form of induction *}
 | 
172  | 
||
173  | 
lemma PlusE [elim!]:  | 
|
174  | 
"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"  | 
|
175  | 
by (auto simp add: Plus_def)  | 
|
| 33961 | 176  | 
|
| 33962 | 177  | 
lemma Plus_eq_empty_conv [simp]: "A <+> B = {} \<longleftrightarrow> A = {} \<and> B = {}"
 | 
178  | 
by auto  | 
|
| 33961 | 179  | 
|
| 33962 | 180  | 
lemma UNIV_Plus_UNIV [simp]: "UNIV <+> UNIV = UNIV"  | 
181  | 
proof (rule set_ext)  | 
|
182  | 
fix u :: "'a + 'b"  | 
|
183  | 
show "u \<in> UNIV <+> UNIV \<longleftrightarrow> u \<in> UNIV" by (cases u) auto  | 
|
184  | 
qed  | 
|
| 33961 | 185  | 
|
| 
36176
 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 
wenzelm 
parents: 
33995 
diff
changeset
 | 
186  | 
hide_const (open) Suml Sumr Projl Projr  | 
| 33961 | 187  | 
|
| 10213 | 188  | 
end  |