| author | huffman | 
| Thu, 03 Jul 2008 18:16:40 +0200 | |
| changeset 27475 | 61b979a2c820 | 
| parent 27104 | 791607529f6d | 
| child 28524 | 644b62cf678f | 
| permissions | -rw-r--r-- | 
| 10213 | 1 | (* Title: HOL/Sum_Type.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1992 University of Cambridge | |
| 5 | *) | |
| 6 | ||
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 7 | header{*The Disjoint Sum of Two Types*}
 | 
| 10213 | 8 | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 9 | theory Sum_Type | 
| 22424 | 10 | imports Typedef Fun | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 11 | begin | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 12 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 13 | text{*The representations of the two injections*}
 | 
| 10213 | 14 | |
| 15 | constdefs | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 16 | Inl_Rep :: "['a, 'a, 'b, bool] => bool" | 
| 10213 | 17 | "Inl_Rep == (%a. %x y p. x=a & p)" | 
| 18 | ||
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 19 | Inr_Rep :: "['b, 'a, 'b, bool] => bool" | 
| 10213 | 20 | "Inr_Rep == (%b. %x y p. y=b & ~p)" | 
| 21 | ||
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 22 | |
| 10213 | 23 | global | 
| 24 | ||
| 25 | typedef (Sum) | |
| 22838 | 26 |   ('a, 'b) "+"          (infixr "+" 10)
 | 
| 10213 | 27 |     = "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}"
 | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 28 | by auto | 
| 10213 | 29 | |
| 30 | local | |
| 31 | ||
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 32 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 33 | text{*abstract constants and syntax*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 34 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 35 | constdefs | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 36 | Inl :: "'a => 'a + 'b" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 37 | "Inl == (%a. Abs_Sum(Inl_Rep(a)))" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 38 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 39 | Inr :: "'b => 'a + 'b" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 40 | "Inr == (%b. Abs_Sum(Inr_Rep(b)))" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 41 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 42 |   Plus :: "['a set, 'b set] => ('a + 'b) set"        (infixr "<+>" 65)
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 43 | "A <+> B == (Inl`A) Un (Inr`B)" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 44 |     --{*disjoint sum for sets; the operator + is overloaded with wrong type!*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 45 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 46 | Part :: "['a set, 'b => 'a] => 'a set" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 47 |    "Part A h == A Int {x. ? z. x = h(z)}"
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 48 |     --{*for selecting out the components of a mutually recursive definition*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 49 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 50 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 51 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 52 | (** Inl_Rep and Inr_Rep: Representations of the constructors **) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 53 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 54 | (*This counts as a non-emptiness result for admitting 'a+'b as a type*) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 55 | lemma Inl_RepI: "Inl_Rep(a) : Sum" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 56 | by (auto simp add: Sum_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 57 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 58 | lemma Inr_RepI: "Inr_Rep(b) : Sum" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 59 | by (auto simp add: Sum_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 60 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 61 | lemma inj_on_Abs_Sum: "inj_on Abs_Sum Sum" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 62 | apply (rule inj_on_inverseI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 63 | apply (erule Abs_Sum_inverse) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 64 | done | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 65 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 66 | subsection{*Freeness Properties for @{term Inl} and  @{term Inr}*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 67 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 68 | text{*Distinctness*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 69 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 70 | lemma Inl_Rep_not_Inr_Rep: "Inl_Rep(a) ~= Inr_Rep(b)" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 71 | by (auto simp add: Inl_Rep_def Inr_Rep_def expand_fun_eq) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 72 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 73 | lemma Inl_not_Inr [iff]: "Inl(a) ~= Inr(b)" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 74 | apply (simp add: Inl_def Inr_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 75 | apply (rule inj_on_Abs_Sum [THEN inj_on_contraD]) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 76 | apply (rule Inl_Rep_not_Inr_Rep) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 77 | apply (rule Inl_RepI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 78 | apply (rule Inr_RepI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 79 | done | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 80 | |
| 17084 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
17026diff
changeset | 81 | lemmas Inr_not_Inl = Inl_not_Inr [THEN not_sym, standard] | 
| 
fb0a80aef0be
classical rules must have names for ATP integration
 paulson parents: 
17026diff
changeset | 82 | declare Inr_not_Inl [iff] | 
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 83 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 84 | lemmas Inl_neq_Inr = Inl_not_Inr [THEN notE, standard] | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 85 | lemmas Inr_neq_Inl = sym [THEN Inl_neq_Inr, standard] | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 86 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 87 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 88 | text{*Injectiveness*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 89 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 90 | lemma Inl_Rep_inject: "Inl_Rep(a) = Inl_Rep(c) ==> a=c" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 91 | by (auto simp add: Inl_Rep_def expand_fun_eq) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 92 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 93 | lemma Inr_Rep_inject: "Inr_Rep(b) = Inr_Rep(d) ==> b=d" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 94 | by (auto simp add: Inr_Rep_def expand_fun_eq) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 95 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 96 | lemma inj_Inl: "inj(Inl)" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 97 | apply (simp add: Inl_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 98 | apply (rule inj_onI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 99 | apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inl_Rep_inject]) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 100 | apply (rule Inl_RepI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 101 | apply (rule Inl_RepI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 102 | done | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 103 | lemmas Inl_inject = inj_Inl [THEN injD, standard] | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 104 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 105 | lemma inj_Inr: "inj(Inr)" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 106 | apply (simp add: Inr_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 107 | apply (rule inj_onI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 108 | apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inr_Rep_inject]) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 109 | apply (rule Inr_RepI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 110 | apply (rule Inr_RepI) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 111 | done | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 112 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 113 | lemmas Inr_inject = inj_Inr [THEN injD, standard] | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 114 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 115 | lemma Inl_eq [iff]: "(Inl(x)=Inl(y)) = (x=y)" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 116 | by (blast dest!: Inl_inject) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 117 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 118 | lemma Inr_eq [iff]: "(Inr(x)=Inr(y)) = (x=y)" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 119 | by (blast dest!: Inr_inject) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 120 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 121 | |
| 22622 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 122 | subsection {* Projections *}
 | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 123 | |
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 124 | definition | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 125 | "sum_case f g x = | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 126 | (if (\<exists>!y. x = Inl y) | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 127 | then f (THE y. x = Inl y) | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 128 | else g (THE y. x = Inr y))" | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 129 | definition "Projl x = sum_case id arbitrary x" | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 130 | definition "Projr x = sum_case arbitrary id x" | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 131 | |
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 132 | lemma sum_cases[simp]: | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 133 | "sum_case f g (Inl x) = f x" | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 134 | "sum_case f g (Inr y) = g y" | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 135 | unfolding sum_case_def | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 136 | by auto | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 137 | |
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 138 | lemma Projl_Inl[simp]: "Projl (Inl x) = x" | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 139 | unfolding Projl_def by simp | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 140 | |
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 141 | lemma Projr_Inr[simp]: "Projr (Inr x) = x" | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 142 | unfolding Projr_def by simp | 
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 143 | |
| 
25693088396b
Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
 krauss parents: 
22424diff
changeset | 144 | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 145 | subsection{*The Disjoint Sum of Sets*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 146 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 147 | (** Introduction rules for the injections **) | 
| 10213 | 148 | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 149 | lemma InlI [intro!]: "a : A ==> Inl(a) : A <+> B" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 150 | by (simp add: Plus_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 151 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 152 | lemma InrI [intro!]: "b : B ==> Inr(b) : A <+> B" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 153 | by (simp add: Plus_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 154 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 155 | (** Elimination rules **) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 156 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 157 | lemma PlusE [elim!]: | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 158 | "[| u: A <+> B; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 159 | !!x. [| x:A; u=Inl(x) |] ==> P; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 160 | !!y. [| y:B; u=Inr(y) |] ==> P | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 161 | |] ==> P" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 162 | by (auto simp add: Plus_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 163 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 164 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 165 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 166 | text{*Exhaustion rule for sums, a degenerate form of induction*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 167 | lemma sumE: | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 168 | "[| !!x::'a. s = Inl(x) ==> P; !!y::'b. s = Inr(y) ==> P | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 169 | |] ==> P" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 170 | apply (rule Abs_Sum_cases [of s]) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 171 | apply (auto simp add: Sum_def Inl_def Inr_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 172 | done | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 173 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 174 | |
| 17026 | 175 | lemma UNIV_Plus_UNIV [simp]: "UNIV <+> UNIV = UNIV" | 
| 176 | apply (rule set_ext) | |
| 177 | apply(rename_tac s) | |
| 178 | apply(rule_tac s=s in sumE) | |
| 179 | apply auto | |
| 180 | done | |
| 181 | ||
| 182 | ||
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 183 | subsection{*The @{term Part} Primitive*}
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 184 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 185 | lemma Part_eqI [intro]: "[| a : A; a=h(b) |] ==> a : Part A h" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 186 | by (auto simp add: Part_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 187 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 188 | lemmas PartI = Part_eqI [OF _ refl, standard] | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 189 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 190 | lemma PartE [elim!]: "[| a : Part A h; !!z. [| a : A; a=h(z) |] ==> P |] ==> P" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 191 | by (auto simp add: Part_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 192 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 193 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 194 | lemma Part_subset: "Part A h <= A" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 195 | by (auto simp add: Part_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 196 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 197 | lemma Part_mono: "A<=B ==> Part A h <= Part B h" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 198 | by blast | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 199 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 200 | lemmas basic_monos = basic_monos Part_mono | 
| 10213 | 201 | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 202 | lemma PartD1: "a : Part A h ==> a : A" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 203 | by (simp add: Part_def) | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 204 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 205 | lemma Part_id: "Part A (%x. x) = A" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 206 | by blast | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 207 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 208 | lemma Part_Int: "Part (A Int B) h = (Part A h) Int (Part B h)" | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 209 | by blast | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 210 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 211 | lemma Part_Collect: "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}"
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 212 | by blast | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 213 | |
| 20588 | 214 | |
| 15391 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 215 | ML | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 216 | {*
 | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 217 | val Inl_RepI = thm "Inl_RepI"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 218 | val Inr_RepI = thm "Inr_RepI"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 219 | val inj_on_Abs_Sum = thm "inj_on_Abs_Sum"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 220 | val Inl_Rep_not_Inr_Rep = thm "Inl_Rep_not_Inr_Rep"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 221 | val Inl_not_Inr = thm "Inl_not_Inr"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 222 | val Inr_not_Inl = thm "Inr_not_Inl"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 223 | val Inl_neq_Inr = thm "Inl_neq_Inr"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 224 | val Inr_neq_Inl = thm "Inr_neq_Inl"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 225 | val Inl_Rep_inject = thm "Inl_Rep_inject"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 226 | val Inr_Rep_inject = thm "Inr_Rep_inject"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 227 | val inj_Inl = thm "inj_Inl"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 228 | val Inl_inject = thm "Inl_inject"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 229 | val inj_Inr = thm "inj_Inr"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 230 | val Inr_inject = thm "Inr_inject"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 231 | val Inl_eq = thm "Inl_eq"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 232 | val Inr_eq = thm "Inr_eq"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 233 | val InlI = thm "InlI"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 234 | val InrI = thm "InrI"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 235 | val PlusE = thm "PlusE"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 236 | val sumE = thm "sumE"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 237 | val Part_eqI = thm "Part_eqI"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 238 | val PartI = thm "PartI"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 239 | val PartE = thm "PartE"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 240 | val Part_subset = thm "Part_subset"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 241 | val Part_mono = thm "Part_mono"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 242 | val PartD1 = thm "PartD1"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 243 | val Part_id = thm "Part_id"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 244 | val Part_Int = thm "Part_Int"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 245 | val Part_Collect = thm "Part_Collect"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 246 | |
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 247 | val basic_monos = thms "basic_monos"; | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 248 | *} | 
| 
797ed46d724b
converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
 paulson parents: 
11609diff
changeset | 249 | |
| 10213 | 250 | end |