src/ZF/Univ.ML
author nipkow
Tue, 17 Oct 2000 08:00:46 +0200
changeset 10228 e653cb933293
parent 9907 473a6604da94
child 10926 27793282952c
permissions -rw-r--r--
added intermediate value thms
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
     1
(*  Title:      ZF/Univ
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents: 187
diff changeset
     4
    Copyright   1994  University of Cambridge
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
The cumulative hierarchy and a small universe for recursive types
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
(*NOT SUITABLE FOR REWRITING -- RECURSIVE!*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
    10
Goal "Vfrom(A,i) = A Un (UN j:i. Pow(Vfrom(A,j)))";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    11
by (stac (Vfrom_def RS def_transrec) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 2033
diff changeset
    12
by (Simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    13
qed "Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
(** Monotonicity **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    17
Goal "A<=B ==> ALL j. i<=j --> Vfrom(A,i) <= Vfrom(B,j)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
by (eps_ind_tac "i" 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
by (rtac (impI RS allI) 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    20
by (stac Vfrom 1);
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    21
by (stac Vfrom 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
by (etac Un_mono 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
by (rtac UN_mono 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    24
by (assume_tac 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
by (rtac Pow_mono 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
by (etac (bspec RS spec RS mp) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
by (assume_tac 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    28
by (rtac subset_refl 1);
3736
39ee3d31cfbc Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents: 3074
diff changeset
    29
qed_spec_mp "Vfrom_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
(** A fundamental equality: Vfrom does not require ordinals! **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
    34
Goal "Vfrom(A,x) <= Vfrom(A,rank(x))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
by (eps_ind_tac "x" 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    36
by (stac Vfrom 1);
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    37
by (stac Vfrom 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
    38
by (blast_tac (claset() addSIs [rank_lt RS ltD]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    39
qed "Vfrom_rank_subset1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
    41
Goal "Vfrom(A,rank(x)) <= Vfrom(A,x)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
by (eps_ind_tac "x" 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    43
by (stac Vfrom 1);
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    44
by (stac Vfrom 1);
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
    45
by (rtac (subset_refl RS Un_mono) 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
    46
by (rtac UN_least 1);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
    47
(*expand rank(x1) = (UN y:x1. succ(rank(y))) in assumptions*)
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
    48
by (etac (rank RS equalityD1 RS subsetD RS UN_E) 1);
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
    49
by (rtac subset_trans 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
    50
by (etac UN_upper 2);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
    51
by (rtac (subset_refl RS Vfrom_mono RS subset_trans RS Pow_mono) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
    52
by (etac (ltI RS le_imp_subset) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
    53
by (rtac (Ord_rank RS Ord_succ) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
by (etac bspec 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    55
by (assume_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    56
qed "Vfrom_rank_subset2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
    58
Goal "Vfrom(A,rank(x)) = Vfrom(A,x)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    59
by (rtac equalityI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    60
by (rtac Vfrom_rank_subset2 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
by (rtac Vfrom_rank_subset1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    62
qed "Vfrom_rank_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    63
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    64
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65
(*** Basic closure properties ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    66
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    67
Goal "y:x ==> 0 : Vfrom(A,x)";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    68
by (stac Vfrom 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2493
diff changeset
    69
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    70
qed "zero_in_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    71
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
    72
Goal "i <= Vfrom(A,i)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    73
by (eps_ind_tac "i" 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    74
by (stac Vfrom 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2493
diff changeset
    75
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    76
qed "i_subset_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    77
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
    78
Goal "A <= Vfrom(A,i)";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    79
by (stac Vfrom 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    80
by (rtac Un_upper1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    81
qed "A_subset_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    82
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
    83
bind_thm ("A_into_Vfrom", A_subset_Vfrom RS subsetD);
488
52f7447d4f1b Addition of infinite branching datatypes
lcp
parents: 484
diff changeset
    84
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    85
Goal "a <= Vfrom(A,i) ==> a: Vfrom(A,succ(i))";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
    86
by (stac Vfrom 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2493
diff changeset
    87
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    88
qed "subset_mem_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    89
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    90
(** Finite sets and ordered pairs **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    91
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
    92
Goal "a: Vfrom(A,i) ==> {a} : Vfrom(A,succ(i))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    93
by (rtac subset_mem_Vfrom 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    94
by Safe_tac;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
    95
qed "singleton_in_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    96
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
    97
Goal "[| a: Vfrom(A,i);  b: Vfrom(A,i) |] ==> {a,b} : Vfrom(A,succ(i))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    98
by (rtac subset_mem_Vfrom 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
    99
by Safe_tac;
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   100
qed "doubleton_in_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   101
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   102
Goalw [Pair_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   103
    "[| a: Vfrom(A,i);  b: Vfrom(A,i) |] ==> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   104
\         <a,b> : Vfrom(A,succ(succ(i)))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   105
by (REPEAT (ares_tac [doubleton_in_Vfrom] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   106
qed "Pair_in_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   107
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   108
Goal "a<=Vfrom(A,i) ==> succ(a) : Vfrom(A,succ(succ(i)))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   109
by (REPEAT (resolve_tac [subset_mem_Vfrom, succ_subsetI] 1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   110
by (rtac (Vfrom_mono RSN (2,subset_trans)) 2);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   111
by (REPEAT (ares_tac [subset_refl, subset_succI] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   112
qed "succ_in_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   113
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   114
(*** 0, successor and limit equations fof Vfrom ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   115
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   116
Goal "Vfrom(A,0) = A";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   117
by (stac Vfrom 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2493
diff changeset
   118
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   119
qed "Vfrom_0";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   120
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   121
Goal "Ord(i) ==> Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   122
by (rtac (Vfrom RS trans) 1);
6
8ce8c4d13d4d Installation of new simplifier for ZF. Deleted all congruence rules not
lcp
parents: 0
diff changeset
   123
by (rtac (succI1 RS RepFunI RS Union_upper RSN
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   124
              (2, equalityI RS subst_context)) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   125
by (rtac UN_least 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   126
by (rtac (subset_refl RS Vfrom_mono RS Pow_mono) 1);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   127
by (etac (ltI RS le_imp_subset) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   128
by (etac Ord_succ 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   129
qed "Vfrom_succ_lemma";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   130
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   131
Goal "Vfrom(A,succ(i)) = A Un Pow(Vfrom(A,i))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   132
by (res_inst_tac [("x1", "succ(i)")] (Vfrom_rank_eq RS subst) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   133
by (res_inst_tac [("x1", "i")] (Vfrom_rank_eq RS subst) 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   134
by (stac rank_succ 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   135
by (rtac (Ord_rank RS Vfrom_succ_lemma) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   136
qed "Vfrom_succ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   137
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   138
(*The premise distinguishes this from Vfrom(A,0);  allowing X=0 forces
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   139
  the conclusion to be Vfrom(A,Union(X)) = A Un (UN y:X. Vfrom(A,y)) *)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   140
Goal "y:X ==> Vfrom(A,Union(X)) = (UN y:X. Vfrom(A,y))";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   141
by (stac Vfrom 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   142
by (rtac equalityI 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   143
(*first inclusion*)
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   144
by (rtac Un_least 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   145
by (rtac (A_subset_Vfrom RS subset_trans) 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   146
by (rtac UN_upper 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   147
by (assume_tac 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   148
by (rtac UN_least 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   149
by (etac UnionE 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   150
by (rtac subset_trans 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   151
by (etac UN_upper 2 THEN stac Vfrom 1 THEN 
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   152
    etac ([UN_upper, Un_upper2] MRS subset_trans) 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   153
(*opposite inclusion*)
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   154
by (rtac UN_least 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   155
by (stac Vfrom 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   156
by (Blast_tac 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   157
qed "Vfrom_Union";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   158
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   159
val [prem] = goal (the_context ()) "y:X ==> Vfrom(A,Union(X)) = (UN y:X. Vfrom(A,y))";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   160
by (stac Vfrom 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   161
by (rtac equalityI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   162
(*first inclusion*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   163
by (rtac Un_least 1);
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   164
by (rtac (A_subset_Vfrom RS subset_trans) 1);
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   165
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   166
by (rtac (prem RS UN_upper) 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   167
by (rtac UN_least 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   168
by (etac UnionE 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   169
by (rtac subset_trans 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   170
by (etac UN_upper 2);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   171
by (stac Vfrom 1);
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   172
by (etac ([UN_upper, Un_upper2] MRS subset_trans) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   173
(*opposite inclusion*)
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   174
by (rtac UN_least 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   175
by (stac Vfrom 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2493
diff changeset
   176
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   177
qed "Vfrom_Union";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   178
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   179
(*** Vfrom applied to Limit ordinals ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   180
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   181
(*NB. limit ordinals are non-empty;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   182
                        Vfrom(A,0) = A = A Un (UN y:0. Vfrom(A,y)) *)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   183
val [limiti] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   184
    "Limit(i) ==> Vfrom(A,i) = (UN y:i. Vfrom(A,y))";
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   185
by (rtac (limiti RS (Limit_has_0 RS ltD) RS Vfrom_Union RS subst) 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   186
by (stac (limiti RS Limit_Union_eq) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   187
by (rtac refl 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   188
qed "Limit_Vfrom_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   189
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   190
Goal "[| a: Vfrom(A,j);  Limit(i);  j<i |] ==> a : Vfrom(A,i)";
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   191
by (rtac (Limit_Vfrom_eq RS equalityD2 RS subsetD) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   192
by (REPEAT (ares_tac [ltD RS UN_I] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   193
qed "Limit_VfromI";
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   194
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   195
val prems = Goal
9173
422968aeed49 fixed some weak elim rules
paulson
parents: 8127
diff changeset
   196
    "[| a: Vfrom(A,i);  ~R ==> Limit(i);               \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   197
\       !!x. [| x<i;  a: Vfrom(A,x) |] ==> R    \
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   198
\    |] ==> R";
9173
422968aeed49 fixed some weak elim rules
paulson
parents: 8127
diff changeset
   199
by (rtac classical 1);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   200
by (rtac (Limit_Vfrom_eq RS equalityD1 RS subsetD RS UN_E) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   201
by (REPEAT (ares_tac (prems @ [ltI, Limit_is_Ord]) 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   202
qed "Limit_VfromE";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   203
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   204
bind_thm ("zero_in_VLimit", Limit_has_0 RS ltD RS zero_in_Vfrom);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   205
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   206
val [major,limiti] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   207
    "[| a: Vfrom(A,i);  Limit(i) |] ==> {a} : Vfrom(A,i)";
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   208
by (rtac ([major,limiti] MRS Limit_VfromE) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   209
by (etac ([singleton_in_Vfrom, limiti] MRS Limit_VfromI) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   210
by (etac (limiti RS Limit_has_succ) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   211
qed "singleton_in_VLimit";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   212
9173
422968aeed49 fixed some weak elim rules
paulson
parents: 8127
diff changeset
   213
bind_thm ("Vfrom_UnI1", Un_upper1 RS (subset_refl RS Vfrom_mono RS subsetD));
422968aeed49 fixed some weak elim rules
paulson
parents: 8127
diff changeset
   214
bind_thm ("Vfrom_UnI2", Un_upper2 RS (subset_refl RS Vfrom_mono RS subsetD));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   215
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   216
(*Hard work is finding a single j:i such that {a,b}<=Vfrom(A,j)*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   217
val [aprem,bprem,limiti] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   218
    "[| a: Vfrom(A,i);  b: Vfrom(A,i);  Limit(i) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   219
\    {a,b} : Vfrom(A,i)";
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   220
by (rtac ([aprem,limiti] MRS Limit_VfromE) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   221
by (rtac ([bprem,limiti] MRS Limit_VfromE) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   222
by (rtac ([doubleton_in_Vfrom, limiti] MRS Limit_VfromI) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   223
by (etac Vfrom_UnI1 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   224
by (etac Vfrom_UnI2 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   225
by (REPEAT (ares_tac [limiti, Limit_has_succ, Un_least_lt] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   226
qed "doubleton_in_VLimit";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   227
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   228
val [aprem,bprem,limiti] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   229
    "[| a: Vfrom(A,i);  b: Vfrom(A,i);  Limit(i) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   230
\    <a,b> : Vfrom(A,i)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   231
(*Infer that a, b occur at ordinals x,xa < i.*)
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   232
by (rtac ([aprem,limiti] MRS Limit_VfromE) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   233
by (rtac ([bprem,limiti] MRS Limit_VfromE) 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   234
by (rtac ([Pair_in_Vfrom, limiti] MRS Limit_VfromI) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   235
(*Infer that succ(succ(x Un xa)) < i *)
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   236
by (etac Vfrom_UnI1 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   237
by (etac Vfrom_UnI2 1);
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   238
by (REPEAT (ares_tac [limiti, Limit_has_succ, Un_least_lt] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   239
qed "Pair_in_VLimit";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   240
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   241
Goal "Limit(i) ==> Vfrom(A,i)*Vfrom(A,i) <= Vfrom(A,i)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   242
by (REPEAT (ares_tac [subsetI,Pair_in_VLimit] 1
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   243
     ORELSE eresolve_tac [SigmaE, ssubst] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   244
qed "product_VLimit";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   245
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   246
bind_thm ("Sigma_subset_VLimit",
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   247
          [Sigma_mono, product_VLimit] MRS subset_trans);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   248
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   249
bind_thm ("nat_subset_VLimit", 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   250
          [nat_le_Limit RS le_imp_subset, i_subset_Vfrom] MRS subset_trans);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   251
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   252
Goal "[| n: nat;  Limit(i) |] ==> n : Vfrom(A,i)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   253
by (REPEAT (ares_tac [nat_subset_VLimit RS subsetD] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   254
qed "nat_into_VLimit";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   255
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   256
(** Closure under disjoint union **)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   257
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   258
bind_thm ("zero_in_VLimit", Limit_has_0 RS ltD RS zero_in_Vfrom);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   259
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   260
Goal "Limit(i) ==> 1 : Vfrom(A,i)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   261
by (REPEAT (ares_tac [nat_into_VLimit, nat_0I, nat_succI] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   262
qed "one_in_VLimit";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   263
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   264
Goalw [Inl_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   265
    "[| a: Vfrom(A,i); Limit(i) |] ==> Inl(a) : Vfrom(A,i)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   266
by (REPEAT (ares_tac [zero_in_VLimit, Pair_in_VLimit] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   267
qed "Inl_in_VLimit";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   268
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   269
Goalw [Inr_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   270
    "[| b: Vfrom(A,i); Limit(i) |] ==> Inr(b) : Vfrom(A,i)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   271
by (REPEAT (ares_tac [one_in_VLimit, Pair_in_VLimit] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   272
qed "Inr_in_VLimit";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   273
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   274
Goal "Limit(i) ==> Vfrom(C,i)+Vfrom(C,i) <= Vfrom(C,i)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   275
by (blast_tac (claset() addSIs [Inl_in_VLimit, Inr_in_VLimit]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   276
qed "sum_VLimit";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   277
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   278
bind_thm ("sum_subset_VLimit", [sum_mono, sum_VLimit] MRS subset_trans);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   279
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   280
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   281
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   282
(*** Properties assuming Transset(A) ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   283
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   284
Goal "Transset(A) ==> Transset(Vfrom(A,i))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   285
by (eps_ind_tac "i" 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   286
by (stac Vfrom 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   287
by (blast_tac (claset() addSIs [Transset_Union_family, Transset_Un,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   288
                            Transset_Pow]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   289
qed "Transset_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   290
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   291
Goal "Transset(A) ==> Vfrom(A, succ(i)) = Pow(Vfrom(A,i))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   292
by (rtac (Vfrom_succ RS trans) 1);
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   293
by (rtac (Un_upper2 RSN (2,equalityI)) 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   294
by (rtac (subset_refl RSN (2,Un_least)) 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   295
by (rtac (A_subset_Vfrom RS subset_trans) 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   296
by (etac (Transset_Vfrom RS (Transset_iff_Pow RS iffD1)) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   297
qed "Transset_Vfrom_succ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   298
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   299
Goalw [Pair_def,Transset_def] "[| <a,b> <= C; Transset(C) |] ==> a: C & b: C";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2493
diff changeset
   300
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   301
qed "Transset_Pair_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   302
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   303
Goal "[| <a,b> <= Vfrom(A,i);  Transset(A);  Limit(i) |] ==> \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   304
\          <a,b> : Vfrom(A,i)";
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   305
by (etac (Transset_Pair_subset RS conjE) 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   306
by (etac Transset_Vfrom 1);
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   307
by (REPEAT (ares_tac [Pair_in_VLimit] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   308
qed "Transset_Pair_subset_VLimit";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   309
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   310
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   311
(*** Closure under product/sum applied to elements -- thus Vfrom(A,i) 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   312
     is a model of simple type theory provided A is a transitive set
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   313
     and i is a limit ordinal
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   314
***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   315
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   316
(*General theorem for membership in Vfrom(A,i) when i is a limit ordinal*)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   317
val [aprem,bprem,limiti,step] = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   318
  "[| a: Vfrom(A,i);  b: Vfrom(A,i);  Limit(i);                 \
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   319
\     !!x y j. [| j<i; 1:j; x: Vfrom(A,j); y: Vfrom(A,j) \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   320
\              |] ==> EX k. h(x,y): Vfrom(A,k) & k<i |] ==>     \
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   321
\  h(a,b) : Vfrom(A,i)";
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   322
(*Infer that a, b occur at ordinals x,xa < i.*)
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   323
by (rtac ([aprem,limiti] MRS Limit_VfromE) 1);
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   324
by (rtac ([bprem,limiti] MRS Limit_VfromE) 1);
828
03d7bfa70289 Changed succ(1) to 2 in in_VLimit, two_in_univ
lcp
parents: 803
diff changeset
   325
by (res_inst_tac [("j1", "x Un xa Un 2")] (step RS exE) 1);
6071
1b2392ac5752 removal of DO_GOAL
paulson
parents: 6053
diff changeset
   326
by (blast_tac (claset() addIs [Limit_VfromI, limiti]) 5);
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   327
by (etac (Vfrom_UnI2 RS Vfrom_UnI1) 4);
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   328
by (etac (Vfrom_UnI1 RS Vfrom_UnI1) 3);
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   329
by (rtac (succI1 RS UnI2) 2);
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   330
by (REPEAT (ares_tac [limiti, Limit_has_0, Limit_has_succ, Un_least_lt] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   331
qed "in_VLimit";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   332
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   333
(** products **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   334
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   335
Goal "[| a: Vfrom(A,j);  b: Vfrom(A,j);  Transset(A) |] ==> \
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   336
\         a*b : Vfrom(A, succ(succ(succ(j))))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   337
by (dtac Transset_Vfrom 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   338
by (rtac subset_mem_Vfrom 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   339
by (rewtac Transset_def);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   340
by (blast_tac (claset() addIs [Pair_in_Vfrom]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   341
qed "prod_in_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   342
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   343
val [aprem,bprem,limiti,transset] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   344
  "[| a: Vfrom(A,i);  b: Vfrom(A,i);  Limit(i);  Transset(A) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   345
\  a*b : Vfrom(A,i)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   346
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1);
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   347
by (REPEAT (ares_tac [exI, conjI, prod_in_Vfrom, transset,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   348
                      limiti RS Limit_has_succ] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   349
qed "prod_in_VLimit";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   350
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   351
(** Disjoint sums, aka Quine ordered pairs **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   352
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   353
Goalw [sum_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   354
    "[| a: Vfrom(A,j);  b: Vfrom(A,j);  Transset(A);  1:j |] ==> \
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   355
\         a+b : Vfrom(A, succ(succ(succ(j))))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   356
by (dtac Transset_Vfrom 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   357
by (rtac subset_mem_Vfrom 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   358
by (rewtac Transset_def);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   359
by (blast_tac (claset() addIs [zero_in_Vfrom, Pair_in_Vfrom, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   360
                           i_subset_Vfrom RS subsetD]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   361
qed "sum_in_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   362
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   363
val [aprem,bprem,limiti,transset] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   364
  "[| a: Vfrom(A,i);  b: Vfrom(A,i);  Limit(i);  Transset(A) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   365
\  a+b : Vfrom(A,i)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   366
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1);
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   367
by (REPEAT (ares_tac [exI, conjI, sum_in_Vfrom, transset,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   368
                      limiti RS Limit_has_succ] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   369
qed "sum_in_VLimit";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   370
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   371
(** function space! **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   372
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   373
Goalw [Pi_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   374
    "[| a: Vfrom(A,j);  b: Vfrom(A,j);  Transset(A) |] ==> \
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   375
\         a->b : Vfrom(A, succ(succ(succ(succ(j)))))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   376
by (dtac Transset_Vfrom 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   377
by (rtac subset_mem_Vfrom 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   378
by (rtac (Collect_subset RS subset_trans) 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   379
by (stac Vfrom 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   380
by (rtac (subset_trans RS subset_trans) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   381
by (rtac Un_upper2 3);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   382
by (rtac (succI1 RS UN_upper) 2);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   383
by (rtac Pow_mono 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   384
by (rewtac Transset_def);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   385
by (blast_tac (claset() addIs [Pair_in_Vfrom]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   386
qed "fun_in_Vfrom";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   387
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   388
val [aprem,bprem,limiti,transset] = goal (the_context ())
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   389
  "[| a: Vfrom(A,i);  b: Vfrom(A,i);  Limit(i);  Transset(A) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   390
\  a->b : Vfrom(A,i)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   391
by (rtac ([aprem,bprem,limiti] MRS in_VLimit) 1);
187
8729bfdcb638 ZF/univ/in_Vfrom_limit: new
lcp
parents: 37
diff changeset
   392
by (REPEAT (ares_tac [exI, conjI, fun_in_Vfrom, transset,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   393
                      limiti RS Limit_has_succ] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   394
qed "fun_in_VLimit";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   395
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   396
Goalw [Pi_def]
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   397
    "[| a: Vfrom(A,j);  Transset(A) |] ==> Pow(a) : Vfrom(A, succ(succ(j)))";
3074
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   398
by (dtac Transset_Vfrom 1);
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   399
by (rtac subset_mem_Vfrom 1);
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   400
by (rewtac Transset_def);
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   401
by (stac Vfrom 1);
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   402
by (Blast_tac 1);
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   403
qed "Pow_in_Vfrom";
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   404
5268
59ef39008514 even more tidying of Goal commands
paulson
parents: 5147
diff changeset
   405
Goal "[| a: Vfrom(A,i);  Limit(i);  Transset(A) |] ==> Pow(a) : Vfrom(A,i)";
5479
5a5dfb0f0d7d fixed PROOF FAILED
paulson
parents: 5321
diff changeset
   406
by (blast_tac (claset() addEs [Limit_VfromE]
9173
422968aeed49 fixed some weak elim rules
paulson
parents: 8127
diff changeset
   407
	                addIs [Limit_has_succ, Pow_in_Vfrom, Limit_VfromI]) 1);
3074
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   408
qed "Pow_in_VLimit";
1fba53dcbf1d New theorems Pow_in_Vfrom and Pow_in_VLimit
paulson
parents: 3016
diff changeset
   409
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   410
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   411
(*** The set Vset(i) ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   412
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   413
Goal "Vset(i) = (UN j:i. Pow(Vset(j)))";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   414
by (stac Vfrom 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2493
diff changeset
   415
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   416
qed "Vset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   417
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   418
bind_thm ("Vset_succ", Transset_0 RS Transset_Vfrom_succ);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   419
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   420
bind_thm ("Transset_Vset", Transset_0 RS Transset_Vfrom);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   421
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   422
(** Characterisation of the elements of Vset(i) **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   423
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   424
Goal "Ord(i) ==> ALL b. b : Vset(i) --> rank(b) < i";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   425
by (etac trans_induct 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   426
by (stac Vset 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   427
by Safe_tac;
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   428
by (stac rank 1);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   429
by (rtac UN_succ_least_lt 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2493
diff changeset
   430
by (Blast_tac 2);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   431
by (REPEAT (ares_tac [ltI] 1));
3736
39ee3d31cfbc Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents: 3074
diff changeset
   432
qed_spec_mp "VsetD";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   433
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   434
Goal "Ord(i) ==> ALL b. rank(b) : i --> b : Vset(i)";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   435
by (etac trans_induct 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   436
by (rtac allI 1);
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   437
by (stac Vset 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   438
by (blast_tac (claset() addSIs [rank_lt RS ltD]) 1);
3736
39ee3d31cfbc Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents: 3074
diff changeset
   439
val lemma = result();
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   440
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   441
Goal "rank(x)<i ==> x : Vset(i)";
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   442
by (etac ltE 1);
3736
39ee3d31cfbc Much tidying including step_tac -> clarify_tac or safe_tac; sometimes
paulson
parents: 3074
diff changeset
   443
by (etac (lemma RS spec RS mp) 1);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   444
by (assume_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   445
qed "VsetI";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   446
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   447
(*Merely a lemma for the result following*)
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   448
Goal "Ord(i) ==> b : Vset(i) <-> rank(b) < i";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   449
by (rtac iffI 1);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   450
by (REPEAT (eresolve_tac [asm_rl, VsetD, VsetI] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   451
qed "Vset_Ord_rank_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   452
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   453
Goal "b : Vset(a) <-> rank(b) < rank(a)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   454
by (rtac (Vfrom_rank_eq RS subst) 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   455
by (rtac (Ord_rank RS Vset_Ord_rank_iff) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   456
qed "Vset_rank_iff";
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   457
Addsimps [Vset_rank_iff];
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   458
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   459
(*This is rank(rank(a)) = rank(a) *)
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   460
Addsimps [Ord_rank RS rank_of_Ord];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   461
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   462
Goal "Ord(i) ==> rank(Vset(i)) = i";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   463
by (stac rank 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   464
by (rtac equalityI 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   465
by Safe_tac;
828
03d7bfa70289 Changed succ(1) to 2 in in_VLimit, two_in_univ
lcp
parents: 803
diff changeset
   466
by (EVERY' [rtac UN_I, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   467
            etac (i_subset_Vfrom RS subsetD),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   468
            etac (Ord_in_Ord RS rank_of_Ord RS ssubst),
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   469
            assume_tac,
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   470
            rtac succI1] 3);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   471
by (REPEAT (eresolve_tac [asm_rl, VsetD RS ltD, Ord_trans] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   472
qed "rank_Vset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   473
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   474
(** Lemmas for reasoning about sets in terms of their elements' ranks **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   475
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   476
Goal "a <= Vset(rank(a))";
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   477
by (rtac subsetI 1);
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   478
by (etac (rank_lt RS VsetI) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   479
qed "arg_subset_Vset_rank";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   480
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   481
val [iprem] = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   482
    "[| !!i. Ord(i) ==> a Int Vset(i) <= b |] ==> a <= b";
27
0e152fe9571e Retry of the previous commit (network outage)
lcp
parents: 15
diff changeset
   483
by (rtac ([subset_refl, arg_subset_Vset_rank] MRS 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   484
          Int_greatest RS subset_trans) 1);
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   485
by (rtac (Ord_rank RS iprem) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   486
qed "Int_Vset_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   487
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   488
(** Set up an environment for simplification **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   489
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   490
Goalw [Inl_def] "rank(a) < rank(Inl(a))";
3889
59bab7a52b4c moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents: 3736
diff changeset
   491
by (rtac rank_pair2 1);
59bab7a52b4c moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents: 3736
diff changeset
   492
qed "rank_Inl";
59bab7a52b4c moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents: 3736
diff changeset
   493
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   494
Goalw [Inr_def] "rank(a) < rank(Inr(a))";
3889
59bab7a52b4c moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents: 3736
diff changeset
   495
by (rtac rank_pair2 1);
59bab7a52b4c moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents: 3736
diff changeset
   496
qed "rank_Inr";
59bab7a52b4c moved rank_Inl, rank_Inr from Epsilon.ML to Univ.ML;
wenzelm
parents: 3736
diff changeset
   497
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   498
bind_thms ("rank_rls", [rank_Inl, rank_Inr, rank_pair1, rank_pair2]);
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   499
bind_thms ("rank_trans_rls", rank_rls @ (rank_rls RLN (2, [lt_trans])));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   500
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   501
val rank_ss = simpset() addsimps [VsetI] addsimps rank_trans_rls;
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   502
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   503
(** Recursion over Vset levels! **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   504
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   505
(*NOT SUITABLE FOR REWRITING: recursive!*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   506
Goalw [Vrec_def] "Vrec(a,H) = H(a, lam x:Vset(rank(a)). Vrec(x,H))";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   507
by (stac transrec 1);
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   508
by (Simp_tac 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   509
by (rtac (refl RS lam_cong RS subst_context) 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   510
by (Asm_full_simp_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   511
qed "Vrec";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   512
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   513
(*This form avoids giant explosions in proofs.  NOTE USE OF == *)
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   514
val rew::prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   515
    "[| !!x. h(x)==Vrec(x,H) |] ==> \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   516
\    h(a) = H(a, lam x: Vset(rank(a)). h(x))";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   517
by (rewtac rew);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   518
by (rtac Vrec 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   519
qed "def_Vrec";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   520
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   521
(*NOT SUITABLE FOR REWRITING: recursive!*)
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   522
Goalw [Vrecursor_def]
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   523
     "Vrecursor(H,a) = H(lam x:Vset(rank(a)). Vrecursor(H,x),  a)";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   524
by (stac transrec 1);
8127
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   525
by (Simp_tac 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   526
by (rtac (refl RS lam_cong RS subst_context) 1);
68c6159440f1 new lemmas for Ntree recursor example; more simprules; more lemmas borrowed
paulson
parents: 6071
diff changeset
   527
by (Asm_full_simp_tac 1);
6053
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   528
qed "Vrecursor";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   529
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   530
(*This form avoids giant explosions in proofs.  NOTE USE OF == *)
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   531
Goal "h == Vrecursor(H) ==> h(a) = H(lam x: Vset(rank(a)). h(x),  a)";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   532
by (Asm_simp_tac 1);
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   533
by (rtac Vrecursor 1);
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   534
qed "def_Vrecursor";
8a1059aa01f0 new inductive, datatype and primrec packages, etc.
paulson
parents: 5479
diff changeset
   535
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   536
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   537
(*** univ(A) ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   538
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   539
Goalw [univ_def] "A<=B ==> univ(A) <= univ(B)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   540
by (etac Vfrom_mono 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   541
by (rtac subset_refl 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   542
qed "univ_mono";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   543
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   544
Goalw [univ_def] "Transset(A) ==> Transset(univ(A))";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   545
by (etac Transset_Vfrom 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   546
qed "Transset_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   547
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   548
(** univ(A) as a limit **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   549
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   550
Goalw [univ_def] "univ(A) = (UN i:nat. Vfrom(A,i))";
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   551
by (rtac (Limit_nat RS Limit_Vfrom_eq) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   552
qed "univ_eq_UN";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   553
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   554
Goal "c <= univ(A) ==> c = (UN i:nat. c Int Vfrom(A,i))";
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   555
by (rtac (subset_UN_iff_eq RS iffD1) 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   556
by (etac (univ_eq_UN RS subst) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   557
qed "subset_univ_eq_Int";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   558
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   559
val [aprem, iprem] = Goal
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   560
    "[| a <= univ(X);                           \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   561
\       !!i. i:nat ==> a Int Vfrom(X,i) <= b    \
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   562
\    |] ==> a <= b";
2033
639de962ded4 Ran expandshort; used stac instead of ssubst
paulson
parents: 1461
diff changeset
   563
by (stac (aprem RS subset_univ_eq_Int) 1);
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   564
by (rtac UN_least 1);
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   565
by (etac iprem 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   566
qed "univ_Int_Vfrom_subset";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   567
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   568
val prems = Goal
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   569
    "[| a <= univ(X);   b <= univ(X);   \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   570
\       !!i. i:nat ==> a Int Vfrom(X,i) = b Int Vfrom(X,i) \
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   571
\    |] ==> a = b";
15
6c6d2f6e3185 ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 6
diff changeset
   572
by (rtac equalityI 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   573
by (ALLGOALS
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   574
    (resolve_tac (prems RL [univ_Int_Vfrom_subset]) THEN'
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   575
     eresolve_tac (prems RL [equalityD1,equalityD2] RL [subset_trans]) THEN'
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   576
     rtac Int_lower1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   577
qed "univ_Int_Vfrom_eq";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   578
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   579
(** Closure properties **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   580
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   581
Goalw [univ_def] "0 : univ(A)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   582
by (rtac (nat_0I RS zero_in_Vfrom) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   583
qed "zero_in_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   584
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   585
Goalw [univ_def] "A <= univ(A)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   586
by (rtac A_subset_Vfrom 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   587
qed "A_subset_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   588
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   589
bind_thm ("A_into_univ", A_subset_univ RS subsetD);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   590
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   591
(** Closure under unordered and ordered pairs **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   592
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   593
Goalw [univ_def] "a: univ(A) ==> {a} : univ(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   594
by (REPEAT (ares_tac [singleton_in_VLimit, Limit_nat] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   595
qed "singleton_in_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   596
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   597
Goalw [univ_def] 
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   598
    "[| a: univ(A);  b: univ(A) |] ==> {a,b} : univ(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   599
by (REPEAT (ares_tac [doubleton_in_VLimit, Limit_nat] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   600
qed "doubleton_in_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   601
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   602
Goalw [univ_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   603
    "[| a: univ(A);  b: univ(A) |] ==> <a,b> : univ(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   604
by (REPEAT (ares_tac [Pair_in_VLimit, Limit_nat] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   605
qed "Pair_in_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   606
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   607
Goalw [univ_def] "univ(A)*univ(A) <= univ(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   608
by (rtac (Limit_nat RS product_VLimit) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   609
qed "product_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   610
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   611
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   612
(** The natural numbers **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   613
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   614
Goalw [univ_def] "nat <= univ(A)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   615
by (rtac i_subset_Vfrom 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   616
qed "nat_subset_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   617
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   618
(* n:nat ==> n:univ(A) *)
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   619
bind_thm ("nat_into_univ", (nat_subset_univ RS subsetD));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   620
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   621
(** instances for 1 and 2 **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   622
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   623
Goalw [univ_def] "1 : univ(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   624
by (rtac (Limit_nat RS one_in_VLimit) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   625
qed "one_in_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   626
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   627
(*unused!*)
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   628
Goal "2 : univ(A)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   629
by (REPEAT (ares_tac [nat_into_univ, nat_0I, nat_succI] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   630
qed "two_in_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   631
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   632
Goalw [bool_def] "bool <= univ(A)";
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   633
by (blast_tac (claset() addSIs [zero_in_univ,one_in_univ]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   634
qed "bool_subset_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   635
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 760
diff changeset
   636
bind_thm ("bool_into_univ", (bool_subset_univ RS subsetD));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   637
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   638
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   639
(** Closure under disjoint union **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   640
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   641
Goalw [univ_def] "a: univ(A) ==> Inl(a) : univ(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   642
by (etac (Limit_nat RSN (2,Inl_in_VLimit)) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   643
qed "Inl_in_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   644
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   645
Goalw [univ_def] "b: univ(A) ==> Inr(b) : univ(A)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   646
by (etac (Limit_nat RSN (2,Inr_in_VLimit)) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   647
qed "Inr_in_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   648
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   649
Goalw [univ_def] "univ(C)+univ(C) <= univ(C)";
516
1957113f0d7d installation of new inductive/datatype sections
lcp
parents: 488
diff changeset
   650
by (rtac (Limit_nat RS sum_VLimit) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 537
diff changeset
   651
qed "sum_univ";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   652
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   653
bind_thm ("sum_subset_univ", [sum_mono, sum_univ] MRS subset_trans);
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   654
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 435
diff changeset
   655
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   656
(** Closure under binary union -- use Un_least **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   657
(** Closure under Collect -- use  (Collect_subset RS subset_trans)  **)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
   658
(** Closure under RepFun -- use   RepFun_subset  **)
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   659
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   660
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   661
(*** Finite Branching Closure Properties ***)
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   662
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   663
(** Closure under finite powerset **)
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   664
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   665
Goal "[| b: Fin(Vfrom(A,i));  Limit(i) |] ==> EX j. b <= Vfrom(A,j) & j<i";
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   666
by (etac Fin_induct 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   667
by (blast_tac (claset() addSDs [Limit_has_0]) 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   668
by Safe_tac;
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   669
by (etac Limit_VfromE 1);
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   670
by (assume_tac 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   671
by (blast_tac (claset() addSIs [Un_least_lt] addIs [Vfrom_UnI1, Vfrom_UnI2]) 1);
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   672
qed "Fin_Vfrom_lemma";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   673
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   674
Goal "Limit(i) ==> Fin(Vfrom(A,i)) <= Vfrom(A,i)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   675
by (rtac subsetI 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   676
by (dtac Fin_Vfrom_lemma 1);
4152
451104c223e2 Ran expandshort, especially to introduce Safe_tac
paulson
parents: 4091
diff changeset
   677
by Safe_tac;
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   678
by (resolve_tac [Vfrom RS ssubst] 1);
4091
771b1f6422a8 isatool fixclasimp;
wenzelm
parents: 3889
diff changeset
   679
by (blast_tac (claset() addSDs [ltD]) 1);
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   680
qed "Fin_VLimit";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   681
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   682
bind_thm ("Fin_subset_VLimit", [Fin_mono, Fin_VLimit] MRS subset_trans);
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   683
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   684
Goalw [univ_def] "Fin(univ(A)) <= univ(A)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   685
by (rtac (Limit_nat RS Fin_VLimit) 1);
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   686
qed "Fin_univ";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   687
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   688
(** Closure under finite powers (functions from a fixed natural number) **)
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   689
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   690
Goal "[| n: nat;  Limit(i) |] ==> n -> Vfrom(A,i) <= Vfrom(A,i)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   691
by (eresolve_tac [nat_fun_subset_Fin RS subset_trans] 1);
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   692
by (REPEAT (ares_tac [Fin_subset_VLimit, Sigma_subset_VLimit,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   693
                      nat_subset_VLimit, subset_refl] 1));
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   694
qed "nat_fun_VLimit";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   695
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   696
bind_thm ("nat_fun_subset_VLimit", [Pi_mono, nat_fun_VLimit] MRS subset_trans);
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   697
5137
60205b0de9b9 Huge tidy-up: removal of leading \!\!
paulson
parents: 5067
diff changeset
   698
Goalw [univ_def] "n: nat ==> n -> univ(A) <= univ(A)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   699
by (etac (Limit_nat RSN (2,nat_fun_VLimit)) 1);
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   700
qed "nat_fun_univ";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   701
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   702
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   703
(** Closure under finite function space **)
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   704
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   705
(*General but seldom-used version; normally the domain is fixed*)
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   706
Goal "Limit(i) ==> Vfrom(A,i) -||> Vfrom(A,i) <= Vfrom(A,i)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   707
by (resolve_tac [FiniteFun.dom_subset RS subset_trans] 1);
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   708
by (REPEAT (ares_tac [Fin_subset_VLimit, Sigma_subset_VLimit, subset_refl] 1));
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   709
qed "FiniteFun_VLimit1";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   710
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   711
Goalw [univ_def] "univ(A) -||> univ(A) <= univ(A)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   712
by (rtac (Limit_nat RS FiniteFun_VLimit1) 1);
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   713
qed "FiniteFun_univ1";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   714
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   715
(*Version for a fixed domain*)
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   716
Goal "[| W <= Vfrom(A,i); Limit(i) |] ==> W -||> Vfrom(A,i) <= Vfrom(A,i)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   717
by (eresolve_tac [subset_refl RSN (2, FiniteFun_mono) RS subset_trans] 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 853
diff changeset
   718
by (etac FiniteFun_VLimit1 1);
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   719
qed "FiniteFun_VLimit";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   720
5067
62b6288e6005 isatool fixgoal;
wenzelm
parents: 4393
diff changeset
   721
Goalw [univ_def]
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   722
    "W <= univ(A) ==> W -||> univ(A) <= univ(A)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   723
by (etac (Limit_nat RSN (2, FiniteFun_VLimit)) 1);
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   724
qed "FiniteFun_univ";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   725
5147
825877190618 More tidying and removal of "\!\!... from Goal commands
paulson
parents: 5137
diff changeset
   726
Goal "[| f: W -||> univ(A);  W <= univ(A) |] ==> f : univ(A)";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   727
by (eresolve_tac [FiniteFun_univ RS subsetD] 1);
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   728
by (assume_tac 1);
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   729
qed "FiniteFun_in_univ";
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   730
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   731
(*Remove <= from the rule above*)
9907
473a6604da94 tuned ML code (the_context, bind_thms(s));
wenzelm
parents: 9173
diff changeset
   732
bind_thm ("FiniteFun_in_univ'", subsetI RSN (2, FiniteFun_in_univ));
803
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   733
4c8333ab3eae changed useless "qed" calls for lemmas back to uses of "result",
lcp
parents: 782
diff changeset
   734
5321
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   735
(**** For QUniv.  Properties of Vfrom analogous to the "take-lemma" ****)
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   736
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   737
(*** Intersecting a*b with Vfrom... ***)
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   738
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   739
(*This version says a, b exist one level down, in the smaller set Vfrom(X,i)*)
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   740
Goal "[| {a,b} : Vfrom(X,succ(i));  Transset(X) |]  \
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   741
\     ==> a: Vfrom(X,i)  &  b: Vfrom(X,i)";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   742
by (dtac (Transset_Vfrom_succ RS equalityD1 RS subsetD RS PowD) 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   743
by (assume_tac 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   744
by (Fast_tac 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   745
qed "doubleton_in_Vfrom_D";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   746
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   747
(*This weaker version says a, b exist at the same level*)
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   748
bind_thm ("Vfrom_doubleton_D", Transset_Vfrom RS Transset_doubleton_D);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   749
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   750
(** Using only the weaker theorem would prove <a,b> : Vfrom(X,i)
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   751
      implies a, b : Vfrom(X,i), which is useless for induction.
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   752
    Using only the stronger theorem would prove <a,b> : Vfrom(X,succ(succ(i)))
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   753
      implies a, b : Vfrom(X,i), leaving the succ(i) case untreated.
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   754
    The combination gives a reduction by precisely one level, which is
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   755
      most convenient for proofs.
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   756
**)
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   757
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   758
Goalw [Pair_def]
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   759
    "[| <a,b> : Vfrom(X,succ(i));  Transset(X) |]  \
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   760
\    ==> a: Vfrom(X,i)  &  b: Vfrom(X,i)";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   761
by (blast_tac (claset() addSDs [doubleton_in_Vfrom_D, Vfrom_doubleton_D]) 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   762
qed "Pair_in_Vfrom_D";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   763
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   764
Goal "Transset(X) ==>          \
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   765
\      (a*b) Int Vfrom(X, succ(i)) <= (a Int Vfrom(X,i)) * (b Int Vfrom(X,i))";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   766
by (blast_tac (claset() addSDs [Pair_in_Vfrom_D]) 1);
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   767
qed "product_Int_Vfrom_subset";
f8848433d240 got rid of some goal thy commands
paulson
parents: 5268
diff changeset
   768