| author | lcp | 
| Thu, 12 Jan 1995 03:02:34 +0100 | |
| changeset 854 | 2e3ca37dfa14 | 
| parent 528 | 61dc99226f8f | 
| child 1155 | 928a16e02f9f | 
| permissions | -rw-r--r-- | 
| 515 | 1 | (* Title: ZF/ex/Brouwer.thy | 
| 2 | ID: $ $ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | Copyright 1994 University of Cambridge | |
| 5 | ||
| 528 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 6 | Infinite branching datatype definitions | 
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 7 | (1) the Brouwer ordinals | 
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 8 | (2) the Martin-Löf wellordering type | 
| 515 | 9 | *) | 
| 10 | ||
| 11 | Brouwer = InfDatatype + | |
| 12 | consts | |
| 13 | brouwer :: "i" | |
| 528 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 14 | Well :: "[i,i=>i]=>i" | 
| 515 | 15 | |
| 16 | datatype <= "Vfrom(0, csucc(nat))" | |
| 17 |   "brouwer" = Zero | Suc ("b: brouwer") | Lim ("h: nat -> brouwer")
 | |
| 18 | monos "[Pi_mono]" | |
| 19 | type_intrs "inf_datatype_intrs" | |
| 20 | ||
| 528 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 21 | (*The union with nat ensures that the cardinal is infinite*) | 
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 22 | datatype <= "Vfrom(A Un (UN x:A. B(x)), csucc(nat Un |UN x:A. B(x)|))" | 
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 23 |   "Well(A,B)" = Sup ("a:A", "f: B(a) -> Well(A,B)")
 | 
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 24 | monos "[Pi_mono]" | 
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 25 | type_intrs "[[UN_upper_cardinal, le_nat_Un_cardinal] MRS le_trans] \ | 
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 26 | \ @ inf_datatype_intrs" | 
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 27 | |
| 
61dc99226f8f
ZF/ex/Brouwer.thy,.ML: new example of wellordering types
 lcp parents: 
515diff
changeset | 28 | |
| 515 | 29 | end |