src/ZF/Univ.thy
author wenzelm
Sun, 25 Oct 2009 19:19:41 +0100
changeset 33170 dd6d8d1f70d2
parent 32960 69916a850301
child 45602 2a858377c3d2
permissions -rw-r--r--
maintain group via name space, not tags; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
     1
(*  Title:      ZF/Univ.thy
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
    Copyright   1992  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
     5
Standard notation for Vset(i) is V(i), but users might want V for a
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
     6
variable.
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 490
diff changeset
     7
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 490
diff changeset
     8
NOTE: univ(A) could be a translation; would simplify many proofs!
6093
87bf8c03b169 eliminated global/local names;
wenzelm
parents: 6053
diff changeset
     9
  But Ind_Syntax.univ refers to the constant "Univ.univ"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
    12
header{*The Cumulative Hierarchy and a Small Universe for Recursive Types*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
    13
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15481
diff changeset
    14
theory Univ imports Epsilon Cardinal begin
3923
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    15
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    16
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    17
  Vfrom       :: "[i,i]=>i"  where
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
    18
    "Vfrom(A,i) == transrec(i, %x f. A Un (\<Union>y\<in>x. Pow(f`y)))"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
24892
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 16417
diff changeset
    20
abbreviation
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 16417
diff changeset
    21
  Vset :: "i=>i" where
c663e675e177 replaced some 'translations' by 'abbreviation';
wenzelm
parents: 16417
diff changeset
    22
  "Vset(x) == Vfrom(0,x)"
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
3923
c257b82a1200 global;
wenzelm
parents: 2469
diff changeset
    24
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    25
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    26
  Vrec        :: "[i, [i,i]=>i] =>i"  where
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    27
    "Vrec(a,H) == transrec(rank(a), %x g. lam z: Vset(succ(x)).
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
    28
                           H(z, lam w:Vset(x). g`rank(w)`w)) ` a"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    29
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    30
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    31
  Vrecursor   :: "[[i,i]=>i, i] =>i"  where
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    32
    "Vrecursor(H,a) == transrec(rank(a), %x g. lam z: Vset(succ(x)).
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 24893
diff changeset
    33
                                H(lam w:Vset(x). g`rank(w)`w, z)) ` a"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    34
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    35
definition
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
    36
  univ        :: "i=>i"  where
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    37
    "univ(A) == Vfrom(A,nat)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    38
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    39
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
    40
subsection{*Immediate Consequences of the Definition of @{term "Vfrom(A,i)"}*}
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
    41
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    42
text{*NOT SUITABLE FOR REWRITING -- RECURSIVE!*}
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
    43
lemma Vfrom: "Vfrom(A,i) = A Un (\<Union>j\<in>i. Pow(Vfrom(A,j)))"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13255
diff changeset
    44
by (subst Vfrom_def [THEN def_transrec], simp)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    45
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    46
subsubsection{* Monotonicity *}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    47
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    48
lemma Vfrom_mono [rule_format]:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
    49
     "A<=B ==> \<forall>j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    50
apply (rule_tac a=i in eps_induct)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    51
apply (rule impI [THEN allI])
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    52
apply (subst Vfrom [of A])
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    53
apply (subst Vfrom [of B])
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    54
apply (erule Un_mono)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    55
apply (erule UN_mono, blast) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    56
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    57
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
    58
lemma VfromI: "[| a \<in> Vfrom(A,j);  j<i |] ==> a \<in> Vfrom(A,i)"
13203
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
    59
by (blast dest: Vfrom_mono [OF subset_refl le_imp_subset [OF leI]]) 
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
    60
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    61
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    62
subsubsection{* A fundamental equality: Vfrom does not require ordinals! *}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    63
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    64
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    65
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    66
lemma Vfrom_rank_subset1: "Vfrom(A,x) <= Vfrom(A,rank(x))"
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    67
proof (induct x rule: eps_induct)
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    68
  fix x
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    69
  assume "\<forall>y\<in>x. Vfrom(A,y) \<subseteq> Vfrom(A,rank(y))"
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    70
  thus "Vfrom(A, x) \<subseteq> Vfrom(A, rank(x))"
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    71
    by (simp add: Vfrom [of _ x] Vfrom [of _ "rank(x)"],
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    72
        blast intro!: rank_lt [THEN ltD])
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    73
qed
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    74
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    75
lemma Vfrom_rank_subset2: "Vfrom(A,rank(x)) <= Vfrom(A,x)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    76
apply (rule_tac a=x in eps_induct)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    77
apply (subst Vfrom)
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 13784
diff changeset
    78
apply (subst Vfrom, rule subset_refl [THEN Un_mono])
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    79
apply (rule UN_least)
13288
9a870391ff66 fixed comment;
wenzelm
parents: 13269
diff changeset
    80
txt{*expand @{text "rank(x1) = (\<Union>y\<in>x1. succ(rank(y)))"} in assumptions*}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    81
apply (erule rank [THEN equalityD1, THEN subsetD, THEN UN_E])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    82
apply (rule subset_trans)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    83
apply (erule_tac [2] UN_upper)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    84
apply (rule subset_refl [THEN Vfrom_mono, THEN subset_trans, THEN Pow_mono])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    85
apply (erule ltI [THEN le_imp_subset])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    86
apply (rule Ord_rank [THEN Ord_succ])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    87
apply (erule bspec, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    88
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    89
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    90
lemma Vfrom_rank_eq: "Vfrom(A,rank(x)) = Vfrom(A,x)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    91
apply (rule equalityI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    92
apply (rule Vfrom_rank_subset2)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    93
apply (rule Vfrom_rank_subset1)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    94
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    95
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    96
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
    97
subsection{* Basic Closure Properties *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
    98
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
    99
lemma zero_in_Vfrom: "y:x ==> 0 \<in> Vfrom(A,x)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   100
by (subst Vfrom, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   101
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   102
lemma i_subset_Vfrom: "i <= Vfrom(A,i)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   103
apply (rule_tac a=i in eps_induct)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   104
apply (subst Vfrom, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   105
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   106
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   107
lemma A_subset_Vfrom: "A <= Vfrom(A,i)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   108
apply (subst Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   109
apply (rule Un_upper1)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   110
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   111
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   112
lemmas A_into_Vfrom = A_subset_Vfrom [THEN subsetD]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   113
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   114
lemma subset_mem_Vfrom: "a <= Vfrom(A,i) ==> a \<in> Vfrom(A,succ(i))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   115
by (subst Vfrom, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   116
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   117
subsubsection{* Finite sets and ordered pairs *}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   118
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   119
lemma singleton_in_Vfrom: "a \<in> Vfrom(A,i) ==> {a} \<in> Vfrom(A,succ(i))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   120
by (rule subset_mem_Vfrom, safe)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   121
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   122
lemma doubleton_in_Vfrom:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   123
     "[| a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i) |] ==> {a,b} \<in> Vfrom(A,succ(i))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   124
by (rule subset_mem_Vfrom, safe)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   125
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   126
lemma Pair_in_Vfrom:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   127
    "[| a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i) |] ==> <a,b> \<in> Vfrom(A,succ(succ(i)))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   128
apply (unfold Pair_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   129
apply (blast intro: doubleton_in_Vfrom) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   130
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   131
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   132
lemma succ_in_Vfrom: "a <= Vfrom(A,i) ==> succ(a) \<in> Vfrom(A,succ(succ(i)))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   133
apply (intro subset_mem_Vfrom succ_subsetI, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   134
apply (erule subset_trans) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   135
apply (rule Vfrom_mono [OF subset_refl subset_succI]) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   136
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   137
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   138
subsection{* 0, Successor and Limit Equations for @{term Vfrom} *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   139
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   140
lemma Vfrom_0: "Vfrom(A,0) = A"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   141
by (subst Vfrom, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   142
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   143
lemma Vfrom_succ_lemma: "Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   144
apply (rule Vfrom [THEN trans])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   145
apply (rule equalityI [THEN subst_context, 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   146
                       OF _ succI1 [THEN RepFunI, THEN Union_upper]])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   147
apply (rule UN_least)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   148
apply (rule subset_refl [THEN Vfrom_mono, THEN Pow_mono])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   149
apply (erule ltI [THEN le_imp_subset])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   150
apply (erule Ord_succ)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   151
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   152
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   153
lemma Vfrom_succ: "Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   154
apply (rule_tac x1 = "succ (i)" in Vfrom_rank_eq [THEN subst])
13784
b9f6154427a4 tidying (by script)
paulson
parents: 13534
diff changeset
   155
apply (rule_tac x1 = i in Vfrom_rank_eq [THEN subst])
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   156
apply (subst rank_succ)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   157
apply (rule Ord_rank [THEN Vfrom_succ_lemma])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   158
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   159
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   160
(*The premise distinguishes this from Vfrom(A,0);  allowing X=0 forces
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   161
  the conclusion to be Vfrom(A,Union(X)) = A Un (\<Union>y\<in>X. Vfrom(A,y)) *)
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   162
lemma Vfrom_Union: "y:X ==> Vfrom(A,Union(X)) = (\<Union>y\<in>X. Vfrom(A,y))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   163
apply (subst Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   164
apply (rule equalityI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   165
txt{*first inclusion*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   166
apply (rule Un_least)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   167
apply (rule A_subset_Vfrom [THEN subset_trans])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   168
apply (rule UN_upper, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   169
apply (rule UN_least)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   170
apply (erule UnionE)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   171
apply (rule subset_trans)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   172
apply (erule_tac [2] UN_upper,
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   173
       subst Vfrom, erule subset_trans [OF UN_upper Un_upper2])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   174
txt{*opposite inclusion*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   175
apply (rule UN_least)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   176
apply (subst Vfrom, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   177
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   178
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   179
subsection{* @{term Vfrom} applied to Limit Ordinals *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   180
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   181
(*NB. limit ordinals are non-empty:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   182
      Vfrom(A,0) = A = A Un (\<Union>y\<in>0. Vfrom(A,y)) *)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   183
lemma Limit_Vfrom_eq:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   184
    "Limit(i) ==> Vfrom(A,i) = (\<Union>y\<in>i. Vfrom(A,y))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   185
apply (rule Limit_has_0 [THEN ltD, THEN Vfrom_Union, THEN subst], assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   186
apply (simp add: Limit_Union_eq) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   187
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   188
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   189
lemma Limit_VfromE:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   190
    "[| a \<in> Vfrom(A,i);  ~R ==> Limit(i);
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   191
        !!x. [| x<i;  a \<in> Vfrom(A,x) |] ==> R
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   192
     |] ==> R"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   193
apply (rule classical)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   194
apply (rule Limit_Vfrom_eq [THEN equalityD1, THEN subsetD, THEN UN_E])
13203
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   195
  prefer 2 apply assumption
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   196
 apply blast 
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   197
apply (blast intro: ltI Limit_is_Ord)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   198
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   199
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   200
lemma singleton_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   201
    "[| a \<in> Vfrom(A,i);  Limit(i) |] ==> {a} \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   202
apply (erule Limit_VfromE, assumption)
13203
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   203
apply (erule singleton_in_Vfrom [THEN VfromI])
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   204
apply (blast intro: Limit_has_succ) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   205
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   206
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   207
lemmas Vfrom_UnI1 = 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   208
    Un_upper1 [THEN subset_refl [THEN Vfrom_mono, THEN subsetD], standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   209
lemmas Vfrom_UnI2 = 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   210
    Un_upper2 [THEN subset_refl [THEN Vfrom_mono, THEN subsetD], standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   211
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   212
text{*Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   213
lemma doubleton_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   214
    "[| a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i) |] ==> {a,b} \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   215
apply (erule Limit_VfromE, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   216
apply (erule Limit_VfromE, assumption)
13203
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   217
apply (blast intro:  VfromI [OF doubleton_in_Vfrom]
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   218
                     Vfrom_UnI1 Vfrom_UnI2 Limit_has_succ Un_least_lt)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   219
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   220
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   221
lemma Pair_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   222
    "[| a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i) |] ==> <a,b> \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   223
txt{*Infer that a, b occur at ordinals x,xa < i.*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   224
apply (erule Limit_VfromE, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   225
apply (erule Limit_VfromE, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   226
txt{*Infer that succ(succ(x Un xa)) < i *}
13203
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   227
apply (blast intro: VfromI [OF Pair_in_Vfrom]
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   228
                    Vfrom_UnI1 Vfrom_UnI2 Limit_has_succ Un_least_lt)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   229
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   230
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   231
lemma product_VLimit: "Limit(i) ==> Vfrom(A,i) * Vfrom(A,i) <= Vfrom(A,i)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   232
by (blast intro: Pair_in_VLimit)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   233
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   234
lemmas Sigma_subset_VLimit =
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   235
     subset_trans [OF Sigma_mono product_VLimit]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   236
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   237
lemmas nat_subset_VLimit =
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   238
     subset_trans [OF nat_le_Limit [THEN le_imp_subset] i_subset_Vfrom]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   239
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   240
lemma nat_into_VLimit: "[| n: nat;  Limit(i) |] ==> n \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   241
by (blast intro: nat_subset_VLimit [THEN subsetD])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   242
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   243
subsubsection{* Closure under Disjoint Union *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   244
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   245
lemmas zero_in_VLimit = Limit_has_0 [THEN ltD, THEN zero_in_Vfrom, standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   246
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   247
lemma one_in_VLimit: "Limit(i) ==> 1 \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   248
by (blast intro: nat_into_VLimit)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   249
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   250
lemma Inl_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   251
    "[| a \<in> Vfrom(A,i); Limit(i) |] ==> Inl(a) \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   252
apply (unfold Inl_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   253
apply (blast intro: zero_in_VLimit Pair_in_VLimit)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   254
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   255
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   256
lemma Inr_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   257
    "[| b \<in> Vfrom(A,i); Limit(i) |] ==> Inr(b) \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   258
apply (unfold Inr_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   259
apply (blast intro: one_in_VLimit Pair_in_VLimit)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   260
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   261
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   262
lemma sum_VLimit: "Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   263
by (blast intro!: Inl_in_VLimit Inr_in_VLimit)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   264
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   265
lemmas sum_subset_VLimit = subset_trans [OF sum_mono sum_VLimit]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   266
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   267
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   268
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   269
subsection{* Properties assuming @{term "Transset(A)"} *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   270
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   271
lemma Transset_Vfrom: "Transset(A) ==> Transset(Vfrom(A,i))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   272
apply (rule_tac a=i in eps_induct)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   273
apply (subst Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   274
apply (blast intro!: Transset_Union_family Transset_Un Transset_Pow)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   275
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   276
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   277
lemma Transset_Vfrom_succ:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   278
     "Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   279
apply (rule Vfrom_succ [THEN trans])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   280
apply (rule equalityI [OF _ Un_upper2])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   281
apply (rule Un_least [OF _ subset_refl])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   282
apply (rule A_subset_Vfrom [THEN subset_trans])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   283
apply (erule Transset_Vfrom [THEN Transset_iff_Pow [THEN iffD1]])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   284
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   285
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   286
lemma Transset_Pair_subset: "[| <a,b> <= C; Transset(C) |] ==> a: C & b: C"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   287
by (unfold Pair_def Transset_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   288
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   289
lemma Transset_Pair_subset_VLimit:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   290
     "[| <a,b> <= Vfrom(A,i);  Transset(A);  Limit(i) |]
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   291
      ==> <a,b> \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   292
apply (erule Transset_Pair_subset [THEN conjE])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   293
apply (erule Transset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   294
apply (blast intro: Pair_in_VLimit)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   295
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   296
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   297
lemma Union_in_Vfrom:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   298
     "[| X \<in> Vfrom(A,j);  Transset(A) |] ==> Union(X) \<in> Vfrom(A, succ(j))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   299
apply (drule Transset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   300
apply (rule subset_mem_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   301
apply (unfold Transset_def, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   302
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   303
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   304
lemma Union_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   305
     "[| X \<in> Vfrom(A,i);  Limit(i);  Transset(A) |] ==> Union(X) \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   306
apply (rule Limit_VfromE, assumption+)
13203
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   307
apply (blast intro: Limit_has_succ VfromI Union_in_Vfrom)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   308
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   309
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   310
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   311
(*** Closure under product/sum applied to elements -- thus Vfrom(A,i)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   312
     is a model of simple type theory provided A is a transitive set
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   313
     and i is a limit ordinal
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   314
***)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   315
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   316
text{*General theorem for membership in Vfrom(A,i) when i is a limit ordinal*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   317
lemma in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   318
  "[| a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i);
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   319
      !!x y j. [| j<i; 1:j; x \<in> Vfrom(A,j); y \<in> Vfrom(A,j) |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   320
               ==> EX k. h(x,y) \<in> Vfrom(A,k) & k<i |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   321
   ==> h(a,b) \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   322
txt{*Infer that a, b occur at ordinals x,xa < i.*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   323
apply (erule Limit_VfromE, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   324
apply (erule Limit_VfromE, assumption, atomize)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   325
apply (drule_tac x=a in spec) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   326
apply (drule_tac x=b in spec) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   327
apply (drule_tac x="x Un xa Un 2" in spec) 
13203
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   328
apply (simp add: Un_least_lt_iff lt_Ord Vfrom_UnI1 Vfrom_UnI2) 
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   329
apply (blast intro: Limit_has_0 Limit_has_succ VfromI)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   330
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   331
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   332
subsubsection{* Products *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   333
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   334
lemma prod_in_Vfrom:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   335
    "[| a \<in> Vfrom(A,j);  b \<in> Vfrom(A,j);  Transset(A) |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   336
     ==> a*b \<in> Vfrom(A, succ(succ(succ(j))))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   337
apply (drule Transset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   338
apply (rule subset_mem_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   339
apply (unfold Transset_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   340
apply (blast intro: Pair_in_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   341
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   342
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   343
lemma prod_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   344
  "[| a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i);  Transset(A) |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   345
   ==> a*b \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   346
apply (erule in_VLimit, assumption+)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   347
apply (blast intro: prod_in_Vfrom Limit_has_succ)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   348
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   349
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   350
subsubsection{* Disjoint Sums, or Quine Ordered Pairs *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   351
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   352
lemma sum_in_Vfrom:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   353
    "[| a \<in> Vfrom(A,j);  b \<in> Vfrom(A,j);  Transset(A);  1:j |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   354
     ==> a+b \<in> Vfrom(A, succ(succ(succ(j))))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   355
apply (unfold sum_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   356
apply (drule Transset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   357
apply (rule subset_mem_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   358
apply (unfold Transset_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   359
apply (blast intro: zero_in_Vfrom Pair_in_Vfrom i_subset_Vfrom [THEN subsetD])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   360
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   361
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   362
lemma sum_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   363
  "[| a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i);  Transset(A) |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   364
   ==> a+b \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   365
apply (erule in_VLimit, assumption+)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   366
apply (blast intro: sum_in_Vfrom Limit_has_succ)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   367
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   368
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   369
subsubsection{* Function Space! *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   370
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   371
lemma fun_in_Vfrom:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   372
    "[| a \<in> Vfrom(A,j);  b \<in> Vfrom(A,j);  Transset(A) |] ==>
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   373
          a->b \<in> Vfrom(A, succ(succ(succ(succ(j)))))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   374
apply (unfold Pi_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   375
apply (drule Transset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   376
apply (rule subset_mem_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   377
apply (rule Collect_subset [THEN subset_trans])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   378
apply (subst Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   379
apply (rule subset_trans [THEN subset_trans])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   380
apply (rule_tac [3] Un_upper2)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   381
apply (rule_tac [2] succI1 [THEN UN_upper])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   382
apply (rule Pow_mono)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   383
apply (unfold Transset_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   384
apply (blast intro: Pair_in_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   385
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   386
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   387
lemma fun_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   388
  "[| a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i);  Transset(A) |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   389
   ==> a->b \<in> Vfrom(A,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   390
apply (erule in_VLimit, assumption+)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   391
apply (blast intro: fun_in_Vfrom Limit_has_succ)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   392
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   393
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   394
lemma Pow_in_Vfrom:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   395
    "[| a \<in> Vfrom(A,j);  Transset(A) |] ==> Pow(a) \<in> Vfrom(A, succ(succ(j)))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   396
apply (drule Transset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   397
apply (rule subset_mem_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   398
apply (unfold Transset_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   399
apply (subst Vfrom, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   400
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   401
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   402
lemma Pow_in_VLimit:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   403
     "[| a \<in> Vfrom(A,i);  Limit(i);  Transset(A) |] ==> Pow(a) \<in> Vfrom(A,i)"
13203
fac77a839aa2 Tidying up. Mainly moving proofs from Main.thy to other (Isar) theory files.
paulson
parents: 13185
diff changeset
   404
by (blast elim: Limit_VfromE intro: Limit_has_succ Pow_in_Vfrom VfromI)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   405
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   406
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   407
subsection{* The Set @{term "Vset(i)"} *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   408
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   409
lemma Vset: "Vset(i) = (\<Union>j\<in>i. Pow(Vset(j)))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   410
by (subst Vfrom, blast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   411
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   412
lemmas Vset_succ = Transset_0 [THEN Transset_Vfrom_succ, standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   413
lemmas Transset_Vset = Transset_0 [THEN Transset_Vfrom, standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   414
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   415
subsubsection{* Characterisation of the elements of @{term "Vset(i)"} *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   416
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   417
lemma VsetD [rule_format]: "Ord(i) ==> \<forall>b. b \<in> Vset(i) --> rank(b) < i"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   418
apply (erule trans_induct)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   419
apply (subst Vset, safe)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   420
apply (subst rank)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   421
apply (blast intro: ltI UN_succ_least_lt) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   422
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   423
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   424
lemma VsetI_lemma [rule_format]:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   425
     "Ord(i) ==> \<forall>b. rank(b) \<in> i --> b \<in> Vset(i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   426
apply (erule trans_induct)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   427
apply (rule allI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   428
apply (subst Vset)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   429
apply (blast intro!: rank_lt [THEN ltD])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   430
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   431
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   432
lemma VsetI: "rank(x)<i ==> x \<in> Vset(i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   433
by (blast intro: VsetI_lemma elim: ltE)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   434
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   435
text{*Merely a lemma for the next result*}
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   436
lemma Vset_Ord_rank_iff: "Ord(i) ==> b \<in> Vset(i) <-> rank(b) < i"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   437
by (blast intro: VsetD VsetI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   438
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   439
lemma Vset_rank_iff [simp]: "b \<in> Vset(a) <-> rank(b) < rank(a)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   440
apply (rule Vfrom_rank_eq [THEN subst])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   441
apply (rule Ord_rank [THEN Vset_Ord_rank_iff])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   442
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   443
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   444
text{*This is rank(rank(a)) = rank(a) *}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   445
declare Ord_rank [THEN rank_of_Ord, simp]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   446
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   447
lemma rank_Vset: "Ord(i) ==> rank(Vset(i)) = i"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   448
apply (subst rank)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   449
apply (rule equalityI, safe)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   450
apply (blast intro: VsetD [THEN ltD]) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   451
apply (blast intro: VsetD [THEN ltD] Ord_trans) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   452
apply (blast intro: i_subset_Vfrom [THEN subsetD]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   453
                    Ord_in_Ord [THEN rank_of_Ord, THEN ssubst])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   454
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   455
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13255
diff changeset
   456
lemma Finite_Vset: "i \<in> nat ==> Finite(Vset(i))";
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13255
diff changeset
   457
apply (erule nat_induct)
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13255
diff changeset
   458
 apply (simp add: Vfrom_0) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13255
diff changeset
   459
apply (simp add: Vset_succ) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13255
diff changeset
   460
done
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13255
diff changeset
   461
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   462
subsubsection{* Reasoning about Sets in Terms of Their Elements' Ranks *}
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   463
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   464
lemma arg_subset_Vset_rank: "a <= Vset(rank(a))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   465
apply (rule subsetI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   466
apply (erule rank_lt [THEN VsetI])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   467
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   468
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   469
lemma Int_Vset_subset:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   470
    "[| !!i. Ord(i) ==> a Int Vset(i) <= b |] ==> a <= b"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   471
apply (rule subset_trans) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   472
apply (rule Int_greatest [OF subset_refl arg_subset_Vset_rank])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   473
apply (blast intro: Ord_rank) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   474
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   475
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   476
subsubsection{* Set Up an Environment for Simplification *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   477
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   478
lemma rank_Inl: "rank(a) < rank(Inl(a))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   479
apply (unfold Inl_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   480
apply (rule rank_pair2)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   481
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   482
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   483
lemma rank_Inr: "rank(a) < rank(Inr(a))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   484
apply (unfold Inr_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   485
apply (rule rank_pair2)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   486
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   487
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   488
lemmas rank_rls = rank_Inl rank_Inr rank_pair1 rank_pair2
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   489
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   490
subsubsection{* Recursion over Vset Levels! *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   491
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   492
text{*NOT SUITABLE FOR REWRITING: recursive!*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   493
lemma Vrec: "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   494
apply (unfold Vrec_def)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13255
diff changeset
   495
apply (subst transrec, simp)
13175
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13163
diff changeset
   496
apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   497
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   498
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   499
text{*This form avoids giant explosions in proofs.  NOTE USE OF == *}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   500
lemma def_Vrec:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   501
    "[| !!x. h(x)==Vrec(x,H) |] ==>
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   502
     h(a) = H(a, lam x: Vset(rank(a)). h(x))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   503
apply simp 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   504
apply (rule Vrec)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   505
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   506
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   507
text{*NOT SUITABLE FOR REWRITING: recursive!*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   508
lemma Vrecursor:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   509
     "Vrecursor(H,a) = H(lam x:Vset(rank(a)). Vrecursor(H,x),  a)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   510
apply (unfold Vrecursor_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   511
apply (subst transrec, simp)
13175
81082cfa5618 new definition of "apply" and new simprule "beta_if"
paulson
parents: 13163
diff changeset
   512
apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   513
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   514
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   515
text{*This form avoids giant explosions in proofs.  NOTE USE OF == *}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   516
lemma def_Vrecursor:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   517
     "h == Vrecursor(H) ==> h(a) = H(lam x: Vset(rank(a)). h(x),  a)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   518
apply simp
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   519
apply (rule Vrecursor)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   520
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   521
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   522
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   523
subsection{* The Datatype Universe: @{term "univ(A)"} *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   524
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   525
lemma univ_mono: "A<=B ==> univ(A) <= univ(B)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   526
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   527
apply (erule Vfrom_mono)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   528
apply (rule subset_refl)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   529
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   530
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   531
lemma Transset_univ: "Transset(A) ==> Transset(univ(A))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   532
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   533
apply (erule Transset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   534
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   535
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   536
subsubsection{* The Set @{term"univ(A)"} as a Limit *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   537
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   538
lemma univ_eq_UN: "univ(A) = (\<Union>i\<in>nat. Vfrom(A,i))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   539
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   540
apply (rule Limit_nat [THEN Limit_Vfrom_eq])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   541
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   542
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   543
lemma subset_univ_eq_Int: "c <= univ(A) ==> c = (\<Union>i\<in>nat. c Int Vfrom(A,i))"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   544
apply (rule subset_UN_iff_eq [THEN iffD1])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   545
apply (erule univ_eq_UN [THEN subst])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   546
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   547
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   548
lemma univ_Int_Vfrom_subset:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   549
    "[| a <= univ(X);
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   550
        !!i. i:nat ==> a Int Vfrom(X,i) <= b |]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   551
     ==> a <= b"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   552
apply (subst subset_univ_eq_Int, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   553
apply (rule UN_least, simp) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   554
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   555
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   556
lemma univ_Int_Vfrom_eq:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   557
    "[| a <= univ(X);   b <= univ(X);
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   558
        !!i. i:nat ==> a Int Vfrom(X,i) = b Int Vfrom(X,i)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   559
     |] ==> a = b"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   560
apply (rule equalityI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   561
apply (rule univ_Int_Vfrom_subset, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   562
apply (blast elim: equalityCE) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   563
apply (rule univ_Int_Vfrom_subset, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   564
apply (blast elim: equalityCE) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   565
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   566
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   567
subsection{* Closure Properties for @{term "univ(A)"}*}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   568
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   569
lemma zero_in_univ: "0 \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   570
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   571
apply (rule nat_0I [THEN zero_in_Vfrom])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   572
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   573
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   574
lemma zero_subset_univ: "{0} <= univ(A)"
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   575
by (blast intro: zero_in_univ)
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   576
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   577
lemma A_subset_univ: "A <= univ(A)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   578
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   579
apply (rule A_subset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   580
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   581
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   582
lemmas A_into_univ = A_subset_univ [THEN subsetD, standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   583
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   584
subsubsection{* Closure under Unordered and Ordered Pairs *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   585
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   586
lemma singleton_in_univ: "a: univ(A) ==> {a} \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   587
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   588
apply (blast intro: singleton_in_VLimit Limit_nat)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   589
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   590
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   591
lemma doubleton_in_univ:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   592
    "[| a: univ(A);  b: univ(A) |] ==> {a,b} \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   593
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   594
apply (blast intro: doubleton_in_VLimit Limit_nat)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   595
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   596
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   597
lemma Pair_in_univ:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   598
    "[| a: univ(A);  b: univ(A) |] ==> <a,b> \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   599
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   600
apply (blast intro: Pair_in_VLimit Limit_nat)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   601
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   602
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   603
lemma Union_in_univ:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   604
     "[| X: univ(A);  Transset(A) |] ==> Union(X) \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   605
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   606
apply (blast intro: Union_in_VLimit Limit_nat)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   607
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   608
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   609
lemma product_univ: "univ(A)*univ(A) <= univ(A)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   610
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   611
apply (rule Limit_nat [THEN product_VLimit])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   612
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   613
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   614
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   615
subsubsection{* The Natural Numbers *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   616
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   617
lemma nat_subset_univ: "nat <= univ(A)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   618
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   619
apply (rule i_subset_Vfrom)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   620
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   621
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   622
text{* n:nat ==> n:univ(A) *}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   623
lemmas nat_into_univ = nat_subset_univ [THEN subsetD, standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   624
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   625
subsubsection{* Instances for 1 and 2 *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   626
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   627
lemma one_in_univ: "1 \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   628
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   629
apply (rule Limit_nat [THEN one_in_VLimit])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   630
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   631
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   632
text{*unused!*}
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   633
lemma two_in_univ: "2 \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   634
by (blast intro: nat_into_univ)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   635
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   636
lemma bool_subset_univ: "bool <= univ(A)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   637
apply (unfold bool_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   638
apply (blast intro!: zero_in_univ one_in_univ)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   639
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   640
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   641
lemmas bool_into_univ = bool_subset_univ [THEN subsetD, standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   642
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   643
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   644
subsubsection{* Closure under Disjoint Union *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   645
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   646
lemma Inl_in_univ: "a: univ(A) ==> Inl(a) \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   647
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   648
apply (erule Inl_in_VLimit [OF _ Limit_nat])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   649
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   650
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   651
lemma Inr_in_univ: "b: univ(A) ==> Inr(b) \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   652
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   653
apply (erule Inr_in_VLimit [OF _ Limit_nat])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   654
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   655
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   656
lemma sum_univ: "univ(C)+univ(C) <= univ(C)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   657
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   658
apply (rule Limit_nat [THEN sum_VLimit])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   659
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   660
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   661
lemmas sum_subset_univ = subset_trans [OF sum_mono sum_univ]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   662
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   663
lemma Sigma_subset_univ:
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   664
  "[|A \<subseteq> univ(D); \<And>x. x \<in> A \<Longrightarrow> B(x) \<subseteq> univ(D)|] ==> Sigma(A,B) \<subseteq> univ(D)"
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   665
apply (simp add: univ_def)
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   666
apply (blast intro: Sigma_subset_VLimit del: subsetI) 
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   667
done
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   668
13255
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   669
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   670
(*Closure under binary union -- use Un_least
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   671
  Closure under Collect -- use  Collect_subset [THEN subset_trans]
407ad9c3036d new theorems, tidying
paulson
parents: 13220
diff changeset
   672
  Closure under RepFun -- use   RepFun_subset *)
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   673
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   674
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   675
subsection{* Finite Branching Closure Properties *}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   676
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   677
subsubsection{* Closure under Finite Powerset *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   678
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   679
lemma Fin_Vfrom_lemma:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   680
     "[| b: Fin(Vfrom(A,i));  Limit(i) |] ==> EX j. b <= Vfrom(A,j) & j<i"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   681
apply (erule Fin_induct)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   682
apply (blast dest!: Limit_has_0, safe)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   683
apply (erule Limit_VfromE, assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   684
apply (blast intro!: Un_least_lt intro: Vfrom_UnI1 Vfrom_UnI2)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   685
done
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   686
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   687
lemma Fin_VLimit: "Limit(i) ==> Fin(Vfrom(A,i)) <= Vfrom(A,i)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   688
apply (rule subsetI)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   689
apply (drule Fin_Vfrom_lemma, safe)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   690
apply (rule Vfrom [THEN ssubst])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   691
apply (blast dest!: ltD)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   692
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   693
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   694
lemmas Fin_subset_VLimit = subset_trans [OF Fin_mono Fin_VLimit]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   695
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   696
lemma Fin_univ: "Fin(univ(A)) <= univ(A)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   697
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   698
apply (rule Limit_nat [THEN Fin_VLimit])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   699
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   700
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   701
subsubsection{* Closure under Finite Powers: Functions from a Natural Number *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   702
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   703
lemma nat_fun_VLimit:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   704
     "[| n: nat;  Limit(i) |] ==> n -> Vfrom(A,i) <= Vfrom(A,i)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   705
apply (erule nat_fun_subset_Fin [THEN subset_trans])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   706
apply (blast del: subsetI
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   707
    intro: subset_refl Fin_subset_VLimit Sigma_subset_VLimit nat_subset_VLimit)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   708
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   709
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   710
lemmas nat_fun_subset_VLimit = subset_trans [OF Pi_mono nat_fun_VLimit]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   711
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   712
lemma nat_fun_univ: "n: nat ==> n -> univ(A) <= univ(A)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   713
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   714
apply (erule nat_fun_VLimit [OF _ Limit_nat])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   715
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   716
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   717
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   718
subsubsection{* Closure under Finite Function Space *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   719
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   720
text{*General but seldom-used version; normally the domain is fixed*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   721
lemma FiniteFun_VLimit1:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   722
     "Limit(i) ==> Vfrom(A,i) -||> Vfrom(A,i) <= Vfrom(A,i)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   723
apply (rule FiniteFun.dom_subset [THEN subset_trans])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   724
apply (blast del: subsetI
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   725
             intro: Fin_subset_VLimit Sigma_subset_VLimit subset_refl)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   726
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   727
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   728
lemma FiniteFun_univ1: "univ(A) -||> univ(A) <= univ(A)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   729
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   730
apply (rule Limit_nat [THEN FiniteFun_VLimit1])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   731
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   732
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   733
text{*Version for a fixed domain*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   734
lemma FiniteFun_VLimit:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   735
     "[| W <= Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) <= Vfrom(A,i)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   736
apply (rule subset_trans) 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   737
apply (erule FiniteFun_mono [OF _ subset_refl])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   738
apply (erule FiniteFun_VLimit1)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   739
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   740
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   741
lemma FiniteFun_univ:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   742
    "W <= univ(A) ==> W -||> univ(A) <= univ(A)"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   743
apply (unfold univ_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   744
apply (erule FiniteFun_VLimit [OF _ Limit_nat])
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   745
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   746
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   747
lemma FiniteFun_in_univ:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   748
     "[| f: W -||> univ(A);  W <= univ(A) |] ==> f \<in> univ(A)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   749
by (erule FiniteFun_univ [THEN subsetD], assumption)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   750
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   751
text{*Remove <= from the rule above*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   752
lemmas FiniteFun_in_univ' = FiniteFun_in_univ [OF _ subsetI]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   753
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   754
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   755
subsection{** For QUniv.  Properties of Vfrom analogous to the "take-lemma" **}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   756
13356
c9cfe1638bf2 improved presentation markup
paulson
parents: 13288
diff changeset
   757
text{* Intersecting a*b with Vfrom... *}
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   758
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   759
text{*This version says a, b exist one level down, in the smaller set Vfrom(X,i)*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   760
lemma doubleton_in_Vfrom_D:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   761
     "[| {a,b} \<in> Vfrom(X,succ(i));  Transset(X) |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   762
      ==> a \<in> Vfrom(X,i)  &  b \<in> Vfrom(X,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   763
by (drule Transset_Vfrom_succ [THEN equalityD1, THEN subsetD, THEN PowD], 
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   764
    assumption, fast)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   765
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   766
text{*This weaker version says a, b exist at the same level*}
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   767
lemmas Vfrom_doubleton_D = Transset_Vfrom [THEN Transset_doubleton_D, standard]
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   768
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   769
(** Using only the weaker theorem would prove <a,b> \<in> Vfrom(X,i)
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   770
      implies a, b \<in> Vfrom(X,i), which is useless for induction.
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   771
    Using only the stronger theorem would prove <a,b> \<in> Vfrom(X,succ(succ(i)))
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   772
      implies a, b \<in> Vfrom(X,i), leaving the succ(i) case untreated.
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   773
    The combination gives a reduction by precisely one level, which is
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   774
      most convenient for proofs.
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   775
**)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   776
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   777
lemma Pair_in_Vfrom_D:
13220
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   778
    "[| <a,b> \<in> Vfrom(X,succ(i));  Transset(X) |]
62c899c77151 tidying
paulson
parents: 13203
diff changeset
   779
     ==> a \<in> Vfrom(X,i)  &  b \<in> Vfrom(X,i)"
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   780
apply (unfold Pair_def)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   781
apply (blast dest!: doubleton_in_Vfrom_D Vfrom_doubleton_D)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   782
done
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   783
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   784
lemma product_Int_Vfrom_subset:
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   785
     "Transset(X) ==>
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   786
      (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))"
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   787
by (blast dest!: Pair_in_Vfrom_D)
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   788
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   789
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   790
ML
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   791
{*
24893
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
   792
val rank_ss = @{simpset} addsimps [@{thm VsetI}] 
b8ef7afe3a6b modernized specifications;
wenzelm
parents: 24892
diff changeset
   793
              addsimps @{thms rank_rls} @ (@{thms rank_rls} RLN (2, [@{thm lt_trans}]));
13163
e320a52ff711 converted Arith, Univ, func to Isar format!
paulson
parents: 9395
diff changeset
   794
*}
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   795
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   796
end