author | bulwahn |
Fri, 18 Mar 2011 18:19:42 +0100 | |
changeset 42031 | 2de57cda5b24 |
parent 39246 | 9e58f0499f57 |
child 46914 | c2ca2c3d23a6 |
permissions | -rw-r--r-- |
11649 | 1 |
(* Title: HOL/Induct/Ordinals.thy |
11641 | 2 |
Author: Stefan Berghofer and Markus Wenzel, TU Muenchen |
3 |
*) |
|
4 |
||
5 |
header {* Ordinals *} |
|
6 |
||
16417 | 7 |
theory Ordinals imports Main begin |
11641 | 8 |
|
9 |
text {* |
|
10 |
Some basic definitions of ordinal numbers. Draws an Agda |
|
11649 | 11 |
development (in Martin-L\"of type theory) by Peter Hancock (see |
11641 | 12 |
\url{http://www.dcs.ed.ac.uk/home/pgh/chat.html}). |
13 |
*} |
|
14 |
||
15 |
datatype ordinal = |
|
16 |
Zero |
|
17 |
| Succ ordinal |
|
18 |
| Limit "nat => ordinal" |
|
19 |
||
39246 | 20 |
primrec pred :: "ordinal => nat => ordinal option" where |
11641 | 21 |
"pred Zero n = None" |
39246 | 22 |
| "pred (Succ a) n = Some a" |
23 |
| "pred (Limit f) n = Some (f n)" |
|
11641 | 24 |
|
39246 | 25 |
abbreviation (input) iter :: "('a => 'a) => nat => ('a => 'a)" where |
26 |
"iter f n \<equiv> f ^^ n" |
|
11641 | 27 |
|
19736 | 28 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
29 |
OpLim :: "(nat => (ordinal => ordinal)) => (ordinal => ordinal)" where |
19736 | 30 |
"OpLim F a = Limit (\<lambda>n. F n a)" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
31 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
32 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
33 |
OpItw :: "(ordinal => ordinal) => (ordinal => ordinal)" ("\<Squnion>") where |
19736 | 34 |
"\<Squnion>f = OpLim (iter f)" |
11641 | 35 |
|
39246 | 36 |
primrec cantor :: "ordinal => ordinal => ordinal" where |
11641 | 37 |
"cantor a Zero = Succ a" |
39246 | 38 |
| "cantor a (Succ b) = \<Squnion>(\<lambda>x. cantor x b) a" |
39 |
| "cantor a (Limit f) = Limit (\<lambda>n. cantor a (f n))" |
|
11641 | 40 |
|
39246 | 41 |
primrec Nabla :: "(ordinal => ordinal) => (ordinal => ordinal)" ("\<nabla>") where |
11641 | 42 |
"\<nabla>f Zero = f Zero" |
39246 | 43 |
| "\<nabla>f (Succ a) = f (Succ (\<nabla>f a))" |
44 |
| "\<nabla>f (Limit h) = Limit (\<lambda>n. \<nabla>f (h n))" |
|
11641 | 45 |
|
19736 | 46 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
47 |
deriv :: "(ordinal => ordinal) => (ordinal => ordinal)" where |
19736 | 48 |
"deriv f = \<nabla>(\<Squnion>f)" |
11641 | 49 |
|
39246 | 50 |
primrec veblen :: "ordinal => ordinal => ordinal" where |
11641 | 51 |
"veblen Zero = \<nabla>(OpLim (iter (cantor Zero)))" |
39246 | 52 |
| "veblen (Succ a) = \<nabla>(OpLim (iter (veblen a)))" |
53 |
| "veblen (Limit f) = \<nabla>(OpLim (\<lambda>n. veblen (f n)))" |
|
11641 | 54 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
55 |
definition "veb a = veblen a Zero" |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
56 |
definition "\<epsilon>\<^isub>0 = veb Zero" |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19736
diff
changeset
|
57 |
definition "\<Gamma>\<^isub>0 = Limit (\<lambda>n. iter veb n Zero)" |
11641 | 58 |
|
59 |
end |