src/ZF/ex/Brouwer.thy
author wenzelm
Sat, 08 Sep 2001 20:02:59 +0200
changeset 11552 aca5398ccd35
parent 11354 9b80fe19407f
permissions -rw-r--r--
ISABELLE_INTERFACE=none by default (cannot expect X11 everywhere);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/ex/Brouwer.thy
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     2
    ID:         $ $
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     5
528
61dc99226f8f ZF/ex/Brouwer.thy,.ML: new example of wellordering types
lcp
parents: 515
diff changeset
     6
Infinite branching datatype definitions
61dc99226f8f ZF/ex/Brouwer.thy,.ML: new example of wellordering types
lcp
parents: 515
diff changeset
     7
  (1) the Brouwer ordinals
61dc99226f8f ZF/ex/Brouwer.thy,.ML: new example of wellordering types
lcp
parents: 515
diff changeset
     8
  (2) the Martin-Löf wellordering type
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
     9
*)
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    10
11354
9b80fe19407f examples files start from Main instead of various ZF theories
paulson
parents: 11316
diff changeset
    11
Brouwer = Main +
9b80fe19407f examples files start from Main instead of various ZF theories
paulson
parents: 11316
diff changeset
    12
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    13
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    14
  brouwer :: i
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    15
  Well    :: [i,i=>i]=>i
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    16
 
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    17
datatype <= "Vfrom(0, csucc(nat))"
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6117
diff changeset
    18
  "brouwer" = Zero | Suc ("b \\<in> brouwer") | Lim ("h \\<in> nat -> brouwer")
6117
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 1478
diff changeset
    19
  monos        Pi_mono
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    20
  type_intrs  "inf_datatype_intrs"
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    21
528
61dc99226f8f ZF/ex/Brouwer.thy,.ML: new example of wellordering types
lcp
parents: 515
diff changeset
    22
(*The union with nat ensures that the cardinal is infinite*)
11316
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6117
diff changeset
    23
datatype <= "Vfrom(A Un (\\<Union>x \\<in> A. B(x)), csucc(nat Un |\\<Union>x \\<in> A. B(x)|))"
b4e71bd751e4 X-symbols for set theory
paulson
parents: 6117
diff changeset
    24
  "Well(A,B)" = Sup ("a \\<in> A", "f \\<in> B(a) -> Well(A,B)")
6117
f9aad8ccd590 tidying of datatype and inductive definitions
paulson
parents: 1478
diff changeset
    25
  monos        Pi_mono
1155
928a16e02f9f removed \...\ inside strings
clasohm
parents: 528
diff changeset
    26
  type_intrs  "[[UN_upper_cardinal, le_nat_Un_cardinal] MRS le_trans]   
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    27
               @ inf_datatype_intrs"
528
61dc99226f8f ZF/ex/Brouwer.thy,.ML: new example of wellordering types
lcp
parents: 515
diff changeset
    28
61dc99226f8f ZF/ex/Brouwer.thy,.ML: new example of wellordering types
lcp
parents: 515
diff changeset
    29
515
abcc438e7c27 installation of new inductive/datatype sections
lcp
parents:
diff changeset
    30
end