| author | clasohm |
| Tue, 24 Oct 1995 13:53:09 +0100 | |
| changeset 1291 | e173be970d27 |
| parent 1211 | 653f33d8d791 |
| child 1461 | 6bcb44e4d6e5 |
| permissions | -rw-r--r-- |
| 992 | 1 |
(* Title: ZF/AC/Transrec2.ML |
2 |
ID: $Id$ |
|
|
1208
bc3093616ba4
Ran expandshort and corrected spelling of Grabczewski
paulson
parents:
992
diff
changeset
|
3 |
Author: Krzysztof Grabczewski |
| 992 | 4 |
|
5 |
Transfinite recursion introduced to handle definitions based on the three |
|
6 |
cases of ordinals. |
|
7 |
*) |
|
8 |
||
9 |
open Transrec2; |
|
10 |
||
11 |
goal thy "transrec2(0,a,b) = a"; |
|
12 |
by (rtac (transrec2_def RS def_transrec RS trans) 1); |
|
13 |
by (simp_tac ZF_ss 1); |
|
14 |
val transrec2_0 = result(); |
|
15 |
||
16 |
goal thy "(THE j. succ(i)=succ(j)) = i"; |
|
| 1211 | 17 |
by (fast_tac (OrdQuant_cs addSIs [the_equality]) 1); |
| 992 | 18 |
val THE_succ_eq = result(); |
19 |
||
20 |
goal thy "transrec2(succ(i),a,b) = b(i, transrec2(i,a,b))"; |
|
21 |
by (rtac (transrec2_def RS def_transrec RS trans) 1); |
|
22 |
by (simp_tac (ZF_ss addsimps [succ_not_0, THE_succ_eq, if_P] |
|
23 |
setsolver K (fast_tac FOL_cs)) 1); |
|
24 |
val transrec2_succ = result(); |
|
25 |
||
26 |
goal thy "!!i. Limit(i) ==> transrec2(i,a,b) = (UN j<i. transrec2(j,a,b))"; |
|
27 |
by (rtac (transrec2_def RS def_transrec RS trans) 1); |
|
28 |
by (resolve_tac [if_not_P RS trans] 1 THEN |
|
| 1211 | 29 |
fast_tac (OrdQuant_cs addSDs [Limit_has_0] addSEs [ltE]) 1); |
| 992 | 30 |
by (resolve_tac [if_not_P RS trans] 1 THEN |
| 1211 | 31 |
fast_tac (OrdQuant_cs addSEs [succ_LimitE]) 1); |
32 |
by (simp_tac OrdQuant_ss 1); |
|
| 992 | 33 |
val transrec2_Limit = result(); |