src/ZF/Sum.thy
author berghofe
Thu, 11 Jul 2002 16:57:14 +0200
changeset 13349 7d4441c8c46a
parent 13255 407ad9c3036d
child 13356 c9cfe1638bf2
permissions -rw-r--r--
Added "using" to the beginning of original newman proof again, because it was lost during last update; renamed second version of newman to newman' (this allows for a comparison of the primitive proof objects, for example).
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/sum.thy
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Disjoint sums in Zermelo-Fraenkel Set Theory 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
"Part" primitive for simultaneous recursive type definitions
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    10
theory Sum = Bool + equalities:
3923
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    11
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    12
global
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    13
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    14
constdefs
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    15
  sum     :: "[i,i]=>i"                     (infixr "+" 65)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    16
     "A+B == {0}*A Un {1}*B"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    17
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    18
  Inl     :: "i=>i"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    19
     "Inl(a) == <0,a>"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    20
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    21
  Inr     :: "i=>i"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    22
     "Inr(b) == <1,b>"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    23
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    24
  "case"  :: "[i=>i, i=>i, i]=>i"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    25
     "case(c,d) == (%<y,z>. cond(y, d(z), c(z)))"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    26
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    27
  (*operator for selecting out the various summands*)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    28
  Part    :: "[i,i=>i] => i"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    29
     "Part(A,h) == {x: A. EX z. x = h(z)}"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
3940
wenzelm
parents: 3923
diff changeset
    31
local
3923
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    32
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    33
(*** Rules for the Part primitive ***)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    34
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    35
lemma Part_iff: 
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    36
    "a : Part(A,h) <-> a:A & (EX y. a=h(y))"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    37
apply (unfold Part_def)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    38
apply (rule separation)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    39
done
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    40
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    41
lemma Part_eqI [intro]: 
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    42
    "[| a : A;  a=h(b) |] ==> a : Part(A,h)"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    43
by (unfold Part_def, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    44
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    45
lemmas PartI = refl [THEN [2] Part_eqI]
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    46
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    47
lemma PartE [elim!]: 
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    48
    "[| a : Part(A,h);  !!z. [| a : A;  a=h(z) |] ==> P   
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    49
     |] ==> P"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    50
apply (unfold Part_def, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    51
done
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    52
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    53
lemma Part_subset: "Part(A,h) <= A"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    54
apply (unfold Part_def)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    55
apply (rule Collect_subset)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    56
done
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    57
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    58
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    59
(*** Rules for Disjoint Sums ***)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    60
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    61
lemmas sum_defs = sum_def Inl_def Inr_def case_def
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    62
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    63
lemma Sigma_bool: "Sigma(bool,C) = C(0) + C(1)"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    64
by (unfold bool_def sum_def, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    65
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    66
(** Introduction rules for the injections **)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    67
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    68
lemma InlI [intro!,simp,TC]: "a : A ==> Inl(a) : A+B"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    69
by (unfold sum_defs, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    70
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    71
lemma InrI [intro!,simp,TC]: "b : B ==> Inr(b) : A+B"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    72
by (unfold sum_defs, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    73
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    74
(** Elimination rules **)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    75
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    76
lemma sumE [elim!]:
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    77
    "[| u: A+B;   
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    78
        !!x. [| x:A;  u=Inl(x) |] ==> P;  
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    79
        !!y. [| y:B;  u=Inr(y) |] ==> P  
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    80
     |] ==> P"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    81
by (unfold sum_defs, blast) 
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    82
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    83
(** Injection and freeness equivalences, for rewriting **)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    84
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    85
lemma Inl_iff [iff]: "Inl(a)=Inl(b) <-> a=b"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    86
by (simp add: sum_defs)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    87
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    88
lemma Inr_iff [iff]: "Inr(a)=Inr(b) <-> a=b"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    89
by (simp add: sum_defs)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    90
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    91
lemma Inl_Inr_iff [iff]: "Inl(a)=Inr(b) <-> False"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    92
by (simp add: sum_defs)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    93
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    94
lemma Inr_Inl_iff [iff]: "Inr(b)=Inl(a) <-> False"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    95
by (simp add: sum_defs)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    96
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    97
lemma sum_empty [simp]: "0+0 = 0"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
    98
by (simp add: sum_defs)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
    99
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   100
(*Injection and freeness rules*)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   101
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   102
lemmas Inl_inject = Inl_iff [THEN iffD1, standard]
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   103
lemmas Inr_inject = Inr_iff [THEN iffD1, standard]
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   104
lemmas Inl_neq_Inr = Inl_Inr_iff [THEN iffD1, THEN FalseE]
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   105
lemmas Inr_neq_Inl = Inr_Inl_iff [THEN iffD1, THEN FalseE]
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   106
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   107
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   108
lemma InlD: "Inl(a): A+B ==> a: A"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   109
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   110
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   111
lemma InrD: "Inr(b): A+B ==> b: B"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   112
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   113
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   114
lemma sum_iff: "u: A+B <-> (EX x. x:A & u=Inl(x)) | (EX y. y:B & u=Inr(y))"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   115
by blast
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   116
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   117
lemma Inl_in_sum_iff [simp]: "(Inl(x) \<in> A+B) <-> (x \<in> A)";
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   118
by auto
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   119
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   120
lemma Inr_in_sum_iff [simp]: "(Inr(y) \<in> A+B) <-> (y \<in> B)";
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   121
by auto
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   122
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   123
lemma sum_subset_iff: "A+B <= C+D <-> A<=C & B<=D"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   124
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   125
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   126
lemma sum_equal_iff: "A+B = C+D <-> A=C & B=D"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   127
by (simp add: extension sum_subset_iff, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   128
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   129
lemma sum_eq_2_times: "A+A = 2*A"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   130
by (simp add: sum_def, blast)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   131
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   132
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   133
(*** Eliminator -- case ***)
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   134
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   135
lemma case_Inl [simp]: "case(c, d, Inl(a)) = c(a)"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   136
by (simp add: sum_defs)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   137
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   138
lemma case_Inr [simp]: "case(c, d, Inr(b)) = d(b)"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   139
by (simp add: sum_defs)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   140
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   141
lemma case_type [TC]:
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   142
    "[| u: A+B;  
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   143
        !!x. x: A ==> c(x): C(Inl(x));    
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   144
        !!y. y: B ==> d(y): C(Inr(y))  
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   145
     |] ==> case(c,d,u) : C(u)"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   146
by auto
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   147
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   148
lemma expand_case: "u: A+B ==>    
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   149
        R(case(c,d,u)) <->  
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   150
        ((ALL x:A. u = Inl(x) --> R(c(x))) &  
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   151
        (ALL y:B. u = Inr(y) --> R(d(y))))"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   152
by auto
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   153
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   154
lemma case_cong:
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   155
  "[| z: A+B;    
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   156
      !!x. x:A ==> c(x)=c'(x);   
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   157
      !!y. y:B ==> d(y)=d'(y)    
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   158
   |] ==> case(c,d,z) = case(c',d',z)"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   159
by auto 
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   160
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   161
lemma case_case: "z: A+B ==>    
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   162
        
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   163
	case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) =  
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   164
        case(%x. c(c'(x)), %y. d(d'(y)), z)"
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   165
by auto
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   166
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   167
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   168
(*** More rules for Part(A,h) ***)
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   169
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   170
lemma Part_mono: "A<=B ==> Part(A,h)<=Part(B,h)"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   171
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   172
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   173
lemma Part_Collect: "Part(Collect(A,P), h) = Collect(Part(A,h), P)"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   174
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   175
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   176
lemmas Part_CollectE =
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   177
     Part_Collect [THEN equalityD1, THEN subsetD, THEN CollectE, standard]
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   178
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   179
lemma Part_Inl: "Part(A+B,Inl) = {Inl(x). x: A}"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   180
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   181
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   182
lemma Part_Inr: "Part(A+B,Inr) = {Inr(y). y: B}"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   183
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   184
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   185
lemma PartD1: "a : Part(A,h) ==> a : A"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   186
by (simp add: Part_def)
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   187
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   188
lemma Part_id: "Part(A,%x. x) = A"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   189
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   190
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   191
lemma Part_Inr2: "Part(A+B, %x. Inr(h(x))) = {Inr(y). y: Part(B,h)}"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   192
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   193
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   194
lemma Part_sum_equality: "C <= A+B ==> Part(C,Inl) Un Part(C,Inr) = C"
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13240
diff changeset
   195
by blast
13240
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   196
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   197
ML
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   198
{*
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   199
val sum_def = thm "sum_def";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   200
val Inl_def = thm "Inl_def";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   201
val Inr_def = thm "Inr_def";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   202
val sum_defs = thms "sum_defs";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   203
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   204
val Part_iff = thm "Part_iff";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   205
val Part_eqI = thm "Part_eqI";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   206
val PartI = thm "PartI";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   207
val PartE = thm "PartE";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   208
val Part_subset = thm "Part_subset";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   209
val Sigma_bool = thm "Sigma_bool";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   210
val InlI = thm "InlI";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   211
val InrI = thm "InrI";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   212
val sumE = thm "sumE";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   213
val Inl_iff = thm "Inl_iff";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   214
val Inr_iff = thm "Inr_iff";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   215
val Inl_Inr_iff = thm "Inl_Inr_iff";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   216
val Inr_Inl_iff = thm "Inr_Inl_iff";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   217
val sum_empty = thm "sum_empty";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   218
val Inl_inject = thm "Inl_inject";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   219
val Inr_inject = thm "Inr_inject";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   220
val Inl_neq_Inr = thm "Inl_neq_Inr";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   221
val Inr_neq_Inl = thm "Inr_neq_Inl";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   222
val InlD = thm "InlD";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   223
val InrD = thm "InrD";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   224
val sum_iff = thm "sum_iff";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   225
val sum_subset_iff = thm "sum_subset_iff";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   226
val sum_equal_iff = thm "sum_equal_iff";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   227
val sum_eq_2_times = thm "sum_eq_2_times";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   228
val case_Inl = thm "case_Inl";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   229
val case_Inr = thm "case_Inr";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   230
val case_type = thm "case_type";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   231
val expand_case = thm "expand_case";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   232
val case_cong = thm "case_cong";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   233
val case_case = thm "case_case";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   234
val Part_mono = thm "Part_mono";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   235
val Part_Collect = thm "Part_Collect";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   236
val Part_CollectE = thm "Part_CollectE";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   237
val Part_Inl = thm "Part_Inl";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   238
val Part_Inr = thm "Part_Inr";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   239
val PartD1 = thm "PartD1";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   240
val Part_id = thm "Part_id";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   241
val Part_Inr2 = thm "Part_Inr2";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   242
val Part_sum_equality = thm "Part_sum_equality";
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   243
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   244
*}
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   245
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   246
bb5f4faea1f3 conversion of Sum, pair to Isar script
paulson
parents: 13220
diff changeset
   247
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   248
end