src/HOL/Sum_Type.thy
changeset 15391 797ed46d724b
parent 11609 3f3d1add4d94
child 17026 43cc86fd3536
equal deleted inserted replaced
15390:87f78411f7c9 15391:797ed46d724b
     1 (*  Title:      HOL/Sum_Type.thy
     1 (*  Title:      HOL/Sum_Type.thy
     2     ID:         $Id$
     2     ID:         $Id$
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     4     Copyright   1992  University of Cambridge
     4     Copyright   1992  University of Cambridge
     5 
       
     6 The disjoint sum of two types.
       
     7 *)
     5 *)
     8 
     6 
     9 Sum_Type = Product_Type +
     7 header{*The Disjoint Sum of Two Types*}
    10 
     8 
    11 (* type definition *)
     9 theory Sum_Type
       
    10 imports Product_Type
       
    11 begin
       
    12 
       
    13 text{*The representations of the two injections*}
    12 
    14 
    13 constdefs
    15 constdefs
    14   Inl_Rep       :: ['a, 'a, 'b, bool] => bool
    16   Inl_Rep :: "['a, 'a, 'b, bool] => bool"
    15   "Inl_Rep == (%a. %x y p. x=a & p)"
    17   "Inl_Rep == (%a. %x y p. x=a & p)"
    16 
    18 
    17   Inr_Rep       :: ['b, 'a, 'b, bool] => bool
    19   Inr_Rep :: "['b, 'a, 'b, bool] => bool"
    18   "Inr_Rep == (%b. %x y p. y=b & ~p)"
    20   "Inr_Rep == (%b. %x y p. y=b & ~p)"
       
    21 
    19 
    22 
    20 global
    23 global
    21 
    24 
    22 typedef (Sum)
    25 typedef (Sum)
    23   ('a, 'b) "+"          (infixr 10)
    26   ('a, 'b) "+"          (infixr 10)
    24     = "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}"
    27     = "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}"
    25 
    28   by auto
    26 
       
    27 (* abstract constants and syntax *)
       
    28 
       
    29 consts
       
    30   Inl            :: "'a => 'a + 'b"
       
    31   Inr            :: "'b => 'a + 'b"
       
    32 
       
    33   (*disjoint sum for sets; the operator + is overloaded with wrong type!*)
       
    34   Plus          :: "['a set, 'b set] => ('a + 'b) set"        (infixr "<+>" 65)
       
    35   Part          :: ['a set, 'b => 'a] => 'a set
       
    36 
    29 
    37 local
    30 local
    38 
    31 
    39 defs
    32 
    40   Inl_def       "Inl == (%a. Abs_Sum(Inl_Rep(a)))"
    33 text{*abstract constants and syntax*}
    41   Inr_def       "Inr == (%b. Abs_Sum(Inr_Rep(b)))"
    34 
    42 
    35 constdefs
    43   sum_def       "A <+> B == (Inl`A) Un (Inr`B)"
    36   Inl :: "'a => 'a + 'b"
    44 
    37    "Inl == (%a. Abs_Sum(Inl_Rep(a)))"
    45   (*for selecting out the components of a mutually recursive definition*)
    38 
    46   Part_def      "Part A h == A Int {x. ? z. x = h(z)}"
    39   Inr :: "'b => 'a + 'b"
       
    40    "Inr == (%b. Abs_Sum(Inr_Rep(b)))"
       
    41 
       
    42   Plus :: "['a set, 'b set] => ('a + 'b) set"        (infixr "<+>" 65)
       
    43    "A <+> B == (Inl`A) Un (Inr`B)"
       
    44     --{*disjoint sum for sets; the operator + is overloaded with wrong type!*}
       
    45 
       
    46   Part :: "['a set, 'b => 'a] => 'a set"
       
    47    "Part A h == A Int {x. ? z. x = h(z)}"
       
    48     --{*for selecting out the components of a mutually recursive definition*}
       
    49 
       
    50 
       
    51 
       
    52 (** Inl_Rep and Inr_Rep: Representations of the constructors **)
       
    53 
       
    54 (*This counts as a non-emptiness result for admitting 'a+'b as a type*)
       
    55 lemma Inl_RepI: "Inl_Rep(a) : Sum"
       
    56 by (auto simp add: Sum_def)
       
    57 
       
    58 lemma Inr_RepI: "Inr_Rep(b) : Sum"
       
    59 by (auto simp add: Sum_def)
       
    60 
       
    61 lemma inj_on_Abs_Sum: "inj_on Abs_Sum Sum"
       
    62 apply (rule inj_on_inverseI)
       
    63 apply (erule Abs_Sum_inverse)
       
    64 done
       
    65 
       
    66 subsection{*Freeness Properties for @{term Inl} and  @{term Inr}*}
       
    67 
       
    68 text{*Distinctness*}
       
    69 
       
    70 lemma Inl_Rep_not_Inr_Rep: "Inl_Rep(a) ~= Inr_Rep(b)"
       
    71 by (auto simp add: Inl_Rep_def Inr_Rep_def expand_fun_eq)
       
    72 
       
    73 lemma Inl_not_Inr [iff]: "Inl(a) ~= Inr(b)"
       
    74 apply (simp add: Inl_def Inr_def)
       
    75 apply (rule inj_on_Abs_Sum [THEN inj_on_contraD])
       
    76 apply (rule Inl_Rep_not_Inr_Rep)
       
    77 apply (rule Inl_RepI)
       
    78 apply (rule Inr_RepI)
       
    79 done
       
    80 
       
    81 lemmas Inr_not_Inl = Inl_not_Inr [THEN not_sym, standard, iff]
       
    82 
       
    83 lemmas Inl_neq_Inr = Inl_not_Inr [THEN notE, standard]
       
    84 lemmas Inr_neq_Inl = sym [THEN Inl_neq_Inr, standard]
       
    85 
       
    86 
       
    87 text{*Injectiveness*}
       
    88 
       
    89 lemma Inl_Rep_inject: "Inl_Rep(a) = Inl_Rep(c) ==> a=c"
       
    90 by (auto simp add: Inl_Rep_def expand_fun_eq)
       
    91 
       
    92 lemma Inr_Rep_inject: "Inr_Rep(b) = Inr_Rep(d) ==> b=d"
       
    93 by (auto simp add: Inr_Rep_def expand_fun_eq)
       
    94 
       
    95 lemma inj_Inl: "inj(Inl)"
       
    96 apply (simp add: Inl_def)
       
    97 apply (rule inj_onI)
       
    98 apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inl_Rep_inject])
       
    99 apply (rule Inl_RepI)
       
   100 apply (rule Inl_RepI)
       
   101 done
       
   102 lemmas Inl_inject = inj_Inl [THEN injD, standard]
       
   103 
       
   104 lemma inj_Inr: "inj(Inr)"
       
   105 apply (simp add: Inr_def)
       
   106 apply (rule inj_onI)
       
   107 apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inr_Rep_inject])
       
   108 apply (rule Inr_RepI)
       
   109 apply (rule Inr_RepI)
       
   110 done
       
   111 
       
   112 lemmas Inr_inject = inj_Inr [THEN injD, standard]
       
   113 
       
   114 lemma Inl_eq [iff]: "(Inl(x)=Inl(y)) = (x=y)"
       
   115 by (blast dest!: Inl_inject)
       
   116 
       
   117 lemma Inr_eq [iff]: "(Inr(x)=Inr(y)) = (x=y)"
       
   118 by (blast dest!: Inr_inject)
       
   119 
       
   120 
       
   121 subsection{*The Disjoint Sum of Sets*}
       
   122 
       
   123 (** Introduction rules for the injections **)
       
   124 
       
   125 lemma InlI [intro!]: "a : A ==> Inl(a) : A <+> B"
       
   126 by (simp add: Plus_def)
       
   127 
       
   128 lemma InrI [intro!]: "b : B ==> Inr(b) : A <+> B"
       
   129 by (simp add: Plus_def)
       
   130 
       
   131 (** Elimination rules **)
       
   132 
       
   133 lemma PlusE [elim!]: 
       
   134     "[| u: A <+> B;   
       
   135         !!x. [| x:A;  u=Inl(x) |] ==> P;  
       
   136         !!y. [| y:B;  u=Inr(y) |] ==> P  
       
   137      |] ==> P"
       
   138 by (auto simp add: Plus_def)
       
   139 
       
   140 
       
   141 
       
   142 text{*Exhaustion rule for sums, a degenerate form of induction*}
       
   143 lemma sumE: 
       
   144     "[| !!x::'a. s = Inl(x) ==> P;  !!y::'b. s = Inr(y) ==> P  
       
   145      |] ==> P"
       
   146 apply (rule Abs_Sum_cases [of s]) 
       
   147 apply (auto simp add: Sum_def Inl_def Inr_def)
       
   148 done
       
   149 
       
   150 lemma sum_induct: "[| !!x. P (Inl x); !!x. P (Inr x) |] ==> P x"
       
   151 by (rule sumE [of x], auto)
       
   152 
       
   153 
       
   154 subsection{*The @{term Part} Primitive*}
       
   155 
       
   156 lemma Part_eqI [intro]: "[| a : A;  a=h(b) |] ==> a : Part A h"
       
   157 by (auto simp add: Part_def)
       
   158 
       
   159 lemmas PartI = Part_eqI [OF _ refl, standard]
       
   160 
       
   161 lemma PartE [elim!]: "[| a : Part A h;  !!z. [| a : A;  a=h(z) |] ==> P |] ==> P"
       
   162 by (auto simp add: Part_def)
       
   163 
       
   164 
       
   165 lemma Part_subset: "Part A h <= A"
       
   166 by (auto simp add: Part_def)
       
   167 
       
   168 lemma Part_mono: "A<=B ==> Part A h <= Part B h"
       
   169 by blast
       
   170 
       
   171 lemmas basic_monos = basic_monos Part_mono
       
   172 
       
   173 lemma PartD1: "a : Part A h ==> a : A"
       
   174 by (simp add: Part_def)
       
   175 
       
   176 lemma Part_id: "Part A (%x. x) = A"
       
   177 by blast
       
   178 
       
   179 lemma Part_Int: "Part (A Int B) h = (Part A h) Int (Part B h)"
       
   180 by blast
       
   181 
       
   182 lemma Part_Collect: "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}"
       
   183 by blast
       
   184 
       
   185 ML
       
   186 {*
       
   187 val Inl_RepI = thm "Inl_RepI";
       
   188 val Inr_RepI = thm "Inr_RepI";
       
   189 val inj_on_Abs_Sum = thm "inj_on_Abs_Sum";
       
   190 val Inl_Rep_not_Inr_Rep = thm "Inl_Rep_not_Inr_Rep";
       
   191 val Inl_not_Inr = thm "Inl_not_Inr";
       
   192 val Inr_not_Inl = thm "Inr_not_Inl";
       
   193 val Inl_neq_Inr = thm "Inl_neq_Inr";
       
   194 val Inr_neq_Inl = thm "Inr_neq_Inl";
       
   195 val Inl_Rep_inject = thm "Inl_Rep_inject";
       
   196 val Inr_Rep_inject = thm "Inr_Rep_inject";
       
   197 val inj_Inl = thm "inj_Inl";
       
   198 val Inl_inject = thm "Inl_inject";
       
   199 val inj_Inr = thm "inj_Inr";
       
   200 val Inr_inject = thm "Inr_inject";
       
   201 val Inl_eq = thm "Inl_eq";
       
   202 val Inr_eq = thm "Inr_eq";
       
   203 val InlI = thm "InlI";
       
   204 val InrI = thm "InrI";
       
   205 val PlusE = thm "PlusE";
       
   206 val sumE = thm "sumE";
       
   207 val sum_induct = thm "sum_induct";
       
   208 val Part_eqI = thm "Part_eqI";
       
   209 val PartI = thm "PartI";
       
   210 val PartE = thm "PartE";
       
   211 val Part_subset = thm "Part_subset";
       
   212 val Part_mono = thm "Part_mono";
       
   213 val PartD1 = thm "PartD1";
       
   214 val Part_id = thm "Part_id";
       
   215 val Part_Int = thm "Part_Int";
       
   216 val Part_Collect = thm "Part_Collect";
       
   217 
       
   218 val basic_monos = thms "basic_monos";
       
   219 *}
       
   220 
    47 
   221 
    48 end
   222 end