| author | paulson | 
| Fri, 14 Jul 2000 14:51:02 +0200 | |
| changeset 9337 | 58bd51302b21 | 
| parent 5068 | fb28eaa07e01 | 
| child 11316 | b4e71bd751e4 | 
| permissions | -rw-r--r-- | 
| 1461 | 1  | 
(* Title: ZF/ex/Brouwer.ML  | 
| 489 | 2  | 
ID: $ $  | 
| 1461 | 3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 489 | 4  | 
Copyright 1994 University of Cambridge  | 
5  | 
||
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
6  | 
Infinite branching datatype definitions  | 
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
7  | 
(1) the Brouwer ordinals  | 
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
8  | 
(2) the Martin-Löf wellordering type  | 
| 489 | 9  | 
*)  | 
10  | 
||
| 515 | 11  | 
open Brouwer;  | 
| 489 | 12  | 
|
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
13  | 
(** The Brouwer ordinals **)  | 
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
14  | 
|
| 5068 | 15  | 
Goal "brouwer = {0} + brouwer + (nat -> brouwer)";
 | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
16  | 
let open brouwer; val rew = rewrite_rule con_defs in  | 
| 4091 | 17  | 
by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)  | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
18  | 
end;  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
525 
diff
changeset
 | 
19  | 
qed "brouwer_unfold";  | 
| 489 | 20  | 
|
21  | 
(*A nicer induction rule than the standard one*)  | 
|
22  | 
val major::prems = goal Brouwer.thy  | 
|
| 1461 | 23  | 
"[| b: brouwer; \  | 
24  | 
\ P(Zero); \  | 
|
25  | 
\ !!b. [| b: brouwer; P(b) |] ==> P(Suc(b)); \  | 
|
26  | 
\ !!h. [| h: nat -> brouwer; ALL i:nat. P(h`i) \  | 
|
27  | 
\ |] ==> P(Lim(h)) \  | 
|
| 489 | 28  | 
\ |] ==> P(b)";  | 
| 515 | 29  | 
by (rtac (major RS brouwer.induct) 1);  | 
| 489 | 30  | 
by (REPEAT_SOME (ares_tac prems));  | 
| 4091 | 31  | 
by (fast_tac (claset() addEs [fun_weaken_type]) 1);  | 
32  | 
by (fast_tac (claset() addDs [apply_type]) 1);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
525 
diff
changeset
 | 
33  | 
qed "brouwer_induct2";  | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
34  | 
|
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
35  | 
|
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
36  | 
(** The Martin-Löf wellordering type **)  | 
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
37  | 
|
| 5068 | 38  | 
Goal "Well(A,B) = (SUM x:A. B(x) -> Well(A,B))";  | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
39  | 
let open Well; val rew = rewrite_rule con_defs in  | 
| 4091 | 40  | 
by (fast_tac (claset() addSIs (map rew intrs) addEs [rew elim]) 1)  | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
41  | 
end;  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
525 
diff
changeset
 | 
42  | 
qed "Well_unfold";  | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
43  | 
|
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
44  | 
(*A nicer induction rule than the standard one*)  | 
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
45  | 
val major::prems = goal Brouwer.thy  | 
| 1461 | 46  | 
"[| w: Well(A,B); \  | 
47  | 
\ !!a f. [| a: A; f: B(a) -> Well(A,B); ALL y: B(a). P(f`y) \  | 
|
48  | 
\ |] ==> P(Sup(a,f)) \  | 
|
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
49  | 
\ |] ==> P(w)";  | 
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
50  | 
by (rtac (major RS Well.induct) 1);  | 
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
51  | 
by (REPEAT_SOME (ares_tac prems));  | 
| 4091 | 52  | 
by (fast_tac (claset() addEs [fun_weaken_type]) 1);  | 
53  | 
by (fast_tac (claset() addDs [apply_type]) 1);  | 
|
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
525 
diff
changeset
 | 
54  | 
qed "Well_induct2";  | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
55  | 
|
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
56  | 
|
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
57  | 
(*In fact it's isomorphic to nat, but we need a recursion operator for  | 
| 
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
58  | 
Well to prove this.*)  | 
| 5068 | 59  | 
Goal "Well(bool, %x. x) = 1 + (1 -> Well(bool, %x. x))";  | 
| 
525
 
e62519a8497d
ZF/List, ex/Brouwer,Data,LList,Ntree,TF,Term: much simplified proof of _unfold
 
lcp 
parents: 
515 
diff
changeset
 | 
60  | 
by (resolve_tac [Well_unfold RS trans] 1);  | 
| 4091 | 61  | 
by (simp_tac (simpset() addsimps [Sigma_bool, Pi_empty1, succ_def]) 1);  | 
| 
782
 
200a16083201
added bind_thm for theorems defined by "standard ..."
 
clasohm 
parents: 
525 
diff
changeset
 | 
62  | 
qed "Well_bool_unfold";  |