src/HOL/Sum_Type.thy
author huffman
Mon, 11 Jun 2007 02:24:39 +0200
changeset 23305 8ae6f7b0903b
parent 22838 466599ecf610
child 25534 d0b74fdd6067
permissions -rw-r--r--
add lemma of_nat_power
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Sum_Type.thy
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     2
    ID:         $Id$
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     4
    Copyright   1992  University of Cambridge
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     5
*)
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     6
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
     7
header{*The Disjoint Sum of Two Types*}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
     8
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
     9
theory Sum_Type
22424
8a5412121687 *** empty log message ***
haftmann
parents: 21454
diff changeset
    10
imports Typedef Fun
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    11
begin
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    12
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    13
text{*The representations of the two injections*}
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    14
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    15
constdefs
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    16
  Inl_Rep :: "['a, 'a, 'b, bool] => bool"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    17
  "Inl_Rep == (%a. %x y p. x=a & p)"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    18
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    19
  Inr_Rep :: "['b, 'a, 'b, bool] => bool"
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    20
  "Inr_Rep == (%b. %x y p. y=b & ~p)"
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    21
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    22
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    23
global
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    24
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    25
typedef (Sum)
22838
haftmann
parents: 22622
diff changeset
    26
  ('a, 'b) "+"          (infixr "+" 10)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    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: 11609
diff changeset
    28
  by auto
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    29
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    30
local
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
    31
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    32
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    33
text{*abstract constants and syntax*}
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    34
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    35
constdefs
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    36
  Inl :: "'a => 'a + 'b"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
    38
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    39
  Inr :: "'b => 'a + 'b"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
    41
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff 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: 11609
diff changeset
    45
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff 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: 11609
diff changeset
    49
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    50
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    51
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
    53
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff changeset
    56
by (auto simp add: Sum_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    57
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
    59
by (auto simp add: Sum_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    60
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
    62
apply (rule inj_on_inverseI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    63
apply (erule Abs_Sum_inverse)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    64
done
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    65
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
    67
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    68
text{*Distinctness*}
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    69
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff changeset
    72
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff 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: 11609
diff 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: 11609
diff changeset
    77
apply (rule Inl_RepI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    78
apply (rule Inr_RepI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    79
done
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    80
17084
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 17026
diff changeset
    81
lemmas Inr_not_Inl = Inl_not_Inr [THEN not_sym, standard]
fb0a80aef0be classical rules must have names for ATP integration
paulson
parents: 17026
diff changeset
    82
declare Inr_not_Inl [iff]
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    83
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff changeset
    86
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    87
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    88
text{*Injectiveness*}
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    89
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff changeset
    92
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff changeset
    95
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    96
lemma inj_Inl: "inj(Inl)"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    97
apply (simp add: Inl_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
    98
apply (rule inj_onI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   100
apply (rule Inl_RepI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   101
apply (rule Inl_RepI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   102
done
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   104
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   105
lemma inj_Inr: "inj(Inr)"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   106
apply (simp add: Inr_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   107
apply (rule inj_onI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   109
apply (rule Inr_RepI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   110
apply (rule Inr_RepI)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   111
done
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   112
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   114
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   116
by (blast dest!: Inl_inject)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   117
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   119
by (blast dest!: Inr_inject)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   120
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   121
22622
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff changeset
   122
subsection {* Projections *}
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff changeset
   123
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff changeset
   124
definition 
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff 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: 22424
diff 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: 22424
diff 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: 22424
diff 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: 22424
diff 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: 22424
diff 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: 22424
diff changeset
   131
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff 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: 22424
diff 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: 22424
diff 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: 22424
diff 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: 22424
diff changeset
   136
  by auto
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff changeset
   137
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff 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: 22424
diff 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: 22424
diff changeset
   140
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff 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: 22424
diff 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: 22424
diff changeset
   143
25693088396b Moving "FunDef" up in the HOL development graph, since it is independent from "Recdef" and "Datatype" now.
krauss
parents: 22424
diff changeset
   144
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   145
subsection{*The Disjoint Sum of Sets*}
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   146
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   147
(** Introduction rules for the injections **)
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   148
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   150
by (simp add: Plus_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   151
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   153
by (simp add: Plus_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   154
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   155
(** Elimination rules **)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   156
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   157
lemma PlusE [elim!]: 
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   158
    "[| u: A <+> B;   
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff changeset
   161
     |] ==> P"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   162
by (auto simp add: Plus_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   163
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   164
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   165
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   167
lemma sumE: 
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff changeset
   169
     |] ==> P"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff 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: 11609
diff 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: 11609
diff changeset
   172
done
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   173
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   174
lemma sum_induct: "[| !!x. P (Inl x); !!x. P (Inr x) |] ==> P x"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   175
by (rule sumE [of x], auto)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   176
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   177
17026
43cc86fd3536 new lemma
nipkow
parents: 15391
diff changeset
   178
lemma UNIV_Plus_UNIV [simp]: "UNIV <+> UNIV = UNIV"
43cc86fd3536 new lemma
nipkow
parents: 15391
diff changeset
   179
apply (rule set_ext)
43cc86fd3536 new lemma
nipkow
parents: 15391
diff changeset
   180
apply(rename_tac s)
43cc86fd3536 new lemma
nipkow
parents: 15391
diff changeset
   181
apply(rule_tac s=s in sumE)
43cc86fd3536 new lemma
nipkow
parents: 15391
diff changeset
   182
apply auto
43cc86fd3536 new lemma
nipkow
parents: 15391
diff changeset
   183
done
43cc86fd3536 new lemma
nipkow
parents: 15391
diff changeset
   184
43cc86fd3536 new lemma
nipkow
parents: 15391
diff changeset
   185
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   186
subsection{*The @{term Part} Primitive*}
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   187
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   188
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: 11609
diff changeset
   189
by (auto simp add: Part_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   190
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   191
lemmas PartI = Part_eqI [OF _ refl, standard]
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   192
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   193
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: 11609
diff changeset
   194
by (auto simp add: Part_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   195
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   196
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   197
lemma Part_subset: "Part A h <= A"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   198
by (auto simp add: Part_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   199
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   200
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: 11609
diff changeset
   201
by blast
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   202
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   203
lemmas basic_monos = basic_monos Part_mono
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   204
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   205
lemma PartD1: "a : Part A h ==> a : A"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   206
by (simp add: Part_def)
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   207
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   208
lemma Part_id: "Part A (%x. x) = A"
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   209
by blast
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   210
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   211
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: 11609
diff changeset
   212
by blast
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   213
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   214
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: 11609
diff changeset
   215
by blast
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   216
20588
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   217
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   218
subsection {* Code generator setup *}
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   219
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   220
instance "+" :: (eq, eq) eq ..
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   221
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   222
lemma [code func]:
21454
a1937c51ed88 dropped eq const
haftmann
parents: 21046
diff changeset
   223
  "(Inl x \<Colon> 'a\<Colon>eq + 'b\<Colon>eq) = Inl y \<longleftrightarrow> x = y"
a1937c51ed88 dropped eq const
haftmann
parents: 21046
diff changeset
   224
  unfolding Inl_eq ..
20588
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   225
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   226
lemma [code func]:
21454
a1937c51ed88 dropped eq const
haftmann
parents: 21046
diff changeset
   227
  "(Inr x \<Colon> 'a\<Colon>eq + 'b\<Colon>eq) = Inr y \<longleftrightarrow> x = y"
a1937c51ed88 dropped eq const
haftmann
parents: 21046
diff changeset
   228
  unfolding Inr_eq ..
20588
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   229
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   230
lemma [code func]:
21454
a1937c51ed88 dropped eq const
haftmann
parents: 21046
diff changeset
   231
  "Inl (x\<Colon>'a\<Colon>eq) = Inr (y\<Colon>'b\<Colon>eq) \<longleftrightarrow> False"
a1937c51ed88 dropped eq const
haftmann
parents: 21046
diff changeset
   232
  using Inl_not_Inr by auto
20588
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   233
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   234
lemma [code func]:
21454
a1937c51ed88 dropped eq const
haftmann
parents: 21046
diff changeset
   235
  "Inr (x\<Colon>'b\<Colon>eq) = Inl (y\<Colon>'a\<Colon>eq) \<longleftrightarrow> False"
a1937c51ed88 dropped eq const
haftmann
parents: 21046
diff changeset
   236
  using Inr_not_Inl by auto
20588
c847c56edf0c added operational equality
haftmann
parents: 20380
diff changeset
   237
15391
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   238
ML
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   239
{*
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   240
val Inl_RepI = thm "Inl_RepI";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   241
val Inr_RepI = thm "Inr_RepI";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   242
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: 11609
diff changeset
   243
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: 11609
diff changeset
   244
val Inl_not_Inr = thm "Inl_not_Inr";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   245
val Inr_not_Inl = thm "Inr_not_Inl";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   246
val Inl_neq_Inr = thm "Inl_neq_Inr";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   247
val Inr_neq_Inl = thm "Inr_neq_Inl";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   248
val Inl_Rep_inject = thm "Inl_Rep_inject";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   249
val Inr_Rep_inject = thm "Inr_Rep_inject";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   250
val inj_Inl = thm "inj_Inl";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   251
val Inl_inject = thm "Inl_inject";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   252
val inj_Inr = thm "inj_Inr";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   253
val Inr_inject = thm "Inr_inject";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   254
val Inl_eq = thm "Inl_eq";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   255
val Inr_eq = thm "Inr_eq";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   256
val InlI = thm "InlI";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   257
val InrI = thm "InrI";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   258
val PlusE = thm "PlusE";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   259
val sumE = thm "sumE";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   260
val sum_induct = thm "sum_induct";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   261
val Part_eqI = thm "Part_eqI";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   262
val PartI = thm "PartI";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   263
val PartE = thm "PartE";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   264
val Part_subset = thm "Part_subset";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   265
val Part_mono = thm "Part_mono";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   266
val PartD1 = thm "PartD1";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   267
val Part_id = thm "Part_id";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   268
val Part_Int = thm "Part_Int";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   269
val Part_Collect = thm "Part_Collect";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   270
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   271
val basic_monos = thms "basic_monos";
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   272
*}
797ed46d724b converted Sum_Type to new-style theory: Inl, Inr are NO LONGER global
paulson
parents: 11609
diff changeset
   273
10213
01c2744a3786 *** empty log message ***
nipkow
parents:
diff changeset
   274
end