src/ZF/QUniv.thy
author wenzelm
Mon, 16 Feb 2009 21:23:33 +0100
changeset 29758 7a3b5bbed313
parent 24893 b8ef7afe3a6b
child 32960 69916a850301
permissions -rw-r--r--
removed rudiments of glossary;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6093
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 3940
diff changeset
     1
(*  Title:      ZF/QUniv.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
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
     8
header{*A Small Universe for Lazy Recursive Types*}
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
     9
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13615
diff changeset
    10
theory QUniv imports Univ QPair begin
3923
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    11
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    12
(*Disjoint sums as a datatype*)
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    13
rep_datatype 
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    14
  elimination	sumE
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    15
  induction	TrueI
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    16
  case_eqns	case_Inl case_Inr
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    17
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    18
(*Variant disjoint sums as a datatype*)
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    19
rep_datatype 
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    20
  elimination	qsumE
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    21
  induction	TrueI
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    22
  case_eqns	qcase_QInl qcase_QInr
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    23
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    24
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 16417
diff changeset
    25
  quniv :: "i => i"  where
6112
5e4871c5136b datatype package improvements
paulson
parents: 6093
diff changeset
    26
   "quniv(A) == Pow(univ(eclose(A)))"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    28
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
    29
subsection{*Properties involving Transset and Sum*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
    30
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    31
lemma Transset_includes_summands:
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    32
     "[| Transset(C); A+B <= C |] ==> A <= C & B <= C"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    33
apply (simp add: sum_def Un_subset_iff) 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    34
apply (blast dest: Transset_includes_range)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    35
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    36
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    37
lemma Transset_sum_Int_subset:
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    38
     "Transset(C) ==> (A+B) Int C <= (A Int C) + (B Int C)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    39
apply (simp add: sum_def Int_Un_distrib2) 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    40
apply (blast dest: Transset_Pair_D)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    41
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    42
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
    43
subsection{*Introduction and Elimination Rules*}
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    44
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    45
lemma qunivI: "X <= univ(eclose(A)) ==> X : quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    46
by (simp add: quniv_def)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    47
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    48
lemma qunivD: "X : quniv(A) ==> X <= univ(eclose(A))"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    49
by (simp add: quniv_def)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    50
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    51
lemma quniv_mono: "A<=B ==> quniv(A) <= quniv(B)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    52
apply (unfold quniv_def)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    53
apply (erule eclose_mono [THEN univ_mono, THEN Pow_mono])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    54
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    55
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
    56
subsection{*Closure Properties*}
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    57
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    58
lemma univ_eclose_subset_quniv: "univ(eclose(A)) <= quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    59
apply (simp add: quniv_def Transset_iff_Pow [symmetric]) 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    60
apply (rule Transset_eclose [THEN Transset_univ])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    61
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    62
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    63
(*Key property for proving A_subset_quniv; requires eclose in def of quniv*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    64
lemma univ_subset_quniv: "univ(A) <= quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    65
apply (rule arg_subset_eclose [THEN univ_mono, THEN subset_trans])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    66
apply (rule univ_eclose_subset_quniv)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    67
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    68
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    69
lemmas univ_into_quniv = univ_subset_quniv [THEN subsetD, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    70
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    71
lemma Pow_univ_subset_quniv: "Pow(univ(A)) <= quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    72
apply (unfold quniv_def)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    73
apply (rule arg_subset_eclose [THEN univ_mono, THEN Pow_mono])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    74
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    75
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    76
lemmas univ_subset_into_quniv =
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    77
    PowI [THEN Pow_univ_subset_quniv [THEN subsetD], standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    78
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    79
lemmas zero_in_quniv = zero_in_univ [THEN univ_into_quniv, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    80
lemmas one_in_quniv = one_in_univ [THEN univ_into_quniv, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    81
lemmas two_in_quniv = two_in_univ [THEN univ_into_quniv, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    82
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    83
lemmas A_subset_quniv =  subset_trans [OF A_subset_univ univ_subset_quniv]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    84
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    85
lemmas A_into_quniv = A_subset_quniv [THEN subsetD, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    86
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    87
(*** univ(A) closure for Quine-inspired pairs and injections ***)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    88
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    89
(*Quine ordered pairs*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    90
lemma QPair_subset_univ: 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    91
    "[| a <= univ(A);  b <= univ(A) |] ==> <a;b> <= univ(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    92
by (simp add: QPair_def sum_subset_univ)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    93
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
    94
subsection{*Quine Disjoint Sum*}
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    95
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    96
lemma QInl_subset_univ: "a <= univ(A) ==> QInl(a) <= univ(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    97
apply (unfold QInl_def)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    98
apply (erule empty_subsetI [THEN QPair_subset_univ])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
    99
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   100
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   101
lemmas naturals_subset_nat = 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   102
    Ord_nat [THEN Ord_is_Transset, unfolded Transset_def, THEN bspec, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   103
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   104
lemmas naturals_subset_univ =
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   105
    subset_trans [OF naturals_subset_nat nat_subset_univ]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   106
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   107
lemma QInr_subset_univ: "a <= univ(A) ==> QInr(a) <= univ(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   108
apply (unfold QInr_def)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   109
apply (erule nat_1I [THEN naturals_subset_univ, THEN QPair_subset_univ])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   110
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   111
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
   112
subsection{*Closure for Quine-Inspired Products and Sums*}
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   113
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   114
(*Quine ordered pairs*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   115
lemma QPair_in_quniv: 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   116
    "[| a: quniv(A);  b: quniv(A) |] ==> <a;b> : quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   117
by (simp add: quniv_def QPair_def sum_subset_univ) 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   118
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   119
lemma QSigma_quniv: "quniv(A) <*> quniv(A) <= quniv(A)" 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   120
by (blast intro: QPair_in_quniv)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   121
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   122
lemmas QSigma_subset_quniv =  subset_trans [OF QSigma_mono QSigma_quniv]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   123
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   124
(*The opposite inclusion*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   125
lemma quniv_QPair_D: 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   126
    "<a;b> : quniv(A) ==> a: quniv(A) & b: quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   127
apply (unfold quniv_def QPair_def)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   128
apply (rule Transset_includes_summands [THEN conjE]) 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   129
apply (rule Transset_eclose [THEN Transset_univ])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   130
apply (erule PowD, blast) 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   131
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   132
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   133
lemmas quniv_QPair_E = quniv_QPair_D [THEN conjE, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   134
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   135
lemma quniv_QPair_iff: "<a;b> : quniv(A) <-> a: quniv(A) & b: quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   136
by (blast intro: QPair_in_quniv dest: quniv_QPair_D)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   137
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   138
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
   139
subsection{*Quine Disjoint Sum*}
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   140
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   141
lemma QInl_in_quniv: "a: quniv(A) ==> QInl(a) : quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   142
by (simp add: QInl_def zero_in_quniv QPair_in_quniv)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   143
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   144
lemma QInr_in_quniv: "b: quniv(A) ==> QInr(b) : quniv(A)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   145
by (simp add: QInr_def one_in_quniv QPair_in_quniv)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   146
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   147
lemma qsum_quniv: "quniv(C) <+> quniv(C) <= quniv(C)"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   148
by (blast intro: QInl_in_quniv QInr_in_quniv)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   149
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   150
lemmas qsum_subset_quniv = subset_trans [OF qsum_mono qsum_quniv]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   151
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   152
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
   153
subsection{*The Natural Numbers*}
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   154
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   155
lemmas nat_subset_quniv =  subset_trans [OF nat_subset_univ univ_subset_quniv]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   156
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   157
(* n:nat ==> n:quniv(A) *)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   158
lemmas nat_into_quniv = nat_subset_quniv [THEN subsetD, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   159
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   160
lemmas bool_subset_quniv = subset_trans [OF bool_subset_univ univ_subset_quniv]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   161
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   162
lemmas bool_into_quniv = bool_subset_quniv [THEN subsetD, standard]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   163
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   164
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
   165
(*Intersecting <a;b> with Vfrom...*)
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   166
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   167
lemma QPair_Int_Vfrom_succ_subset: 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   168
 "Transset(X) ==>           
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   169
       <a;b> Int Vfrom(X, succ(i))  <=  <a Int Vfrom(X,i);  b Int Vfrom(X,i)>"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   170
by (simp add: QPair_def sum_def Int_Un_distrib2 Un_mono
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   171
              product_Int_Vfrom_subset [THEN subset_trans]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   172
              Sigma_mono [OF Int_lower1 subset_refl])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   173
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
   174
subsection{*"Take-Lemma" Rules*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
   175
c9cfe1638bf2 improved presentation markup
paulson
parents: 13285
diff changeset
   176
(*for proving a=b by coinduction and c: quniv(A)*)
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   177
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   178
(*Rule for level i -- preserving the level, not decreasing it*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   179
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   180
lemma QPair_Int_Vfrom_subset: 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   181
 "Transset(X) ==>           
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   182
       <a;b> Int Vfrom(X,i)  <=  <a Int Vfrom(X,i);  b Int Vfrom(X,i)>"
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   183
apply (unfold QPair_def)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   184
apply (erule Transset_Vfrom [THEN Transset_sum_Int_subset])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   185
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   186
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   187
(*[| a Int Vset(i) <= c; b Int Vset(i) <= d |] ==> <a;b> Int Vset(i) <= <c;d>*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   188
lemmas QPair_Int_Vset_subset_trans =
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   189
     subset_trans [OF Transset_0 [THEN QPair_Int_Vfrom_subset] QPair_mono]
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   190
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   191
lemma QPair_Int_Vset_subset_UN:
13615
449a70d88b38 Numerous cosmetic changes, prompted by the new simplifier
paulson
parents: 13357
diff changeset
   192
     "Ord(i) ==> <a;b> Int Vset(i) <= (\<Union>j\<in>i. <a Int Vset(j); b Int Vset(j)>)"
13285
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   193
apply (erule Ord_cases)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   194
(*0 case*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   195
apply (simp add: Vfrom_0)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   196
(*succ(j) case*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   197
apply (erule ssubst) 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   198
apply (rule Transset_0 [THEN QPair_Int_Vfrom_succ_subset, THEN subset_trans])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   199
apply (rule succI1 [THEN UN_upper])
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   200
(*Limit(i) case*)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   201
apply (simp del: UN_simps 
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   202
        add: Limit_Vfrom_eq Int_UN_distrib UN_mono QPair_Int_Vset_subset_trans)
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   203
done
28d1823ce0f2 conversion of QPair to Isar
paulson
parents: 13220
diff changeset
   204
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   205
end