| author | wenzelm | 
| Fri, 29 Oct 2010 11:07:21 +0200 | |
| changeset 40253 | f99ec71de82d | 
| parent 35437 | fe196f61b970 | 
| permissions | -rw-r--r-- | 
| 16417 | 1 | theory Ordinal imports Main begin | 
| 13262 | 2 | |
| 3 | datatype ordinal = Zero | Succ ordinal | Limit "nat \<Rightarrow> ordinal" | |
| 4 | ||
| 5 | consts | |
| 6 | pred :: "ordinal \<Rightarrow> nat \<Rightarrow> ordinal option" | |
| 7 | primrec | |
| 8 | "pred Zero n = None" | |
| 9 | "pred (Succ a) n = Some a" | |
| 10 | "pred (Limit f) n = Some (f n)" | |
| 11 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 12 | definition OpLim :: "(nat \<Rightarrow> (ordinal \<Rightarrow> ordinal)) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" where | 
| 13262 | 13 | "OpLim F a \<equiv> Limit (\<lambda>n. F n a)" | 
| 35437 | 14 | |
| 15 | definition OpItw :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"    ("\<Squnion>") where
 | |
| 13262 | 16 | "\<Squnion>f \<equiv> OpLim (power f)" | 
| 17 | ||
| 18 | consts | |
| 19 | cantor :: "ordinal \<Rightarrow> ordinal \<Rightarrow> ordinal" | |
| 20 | primrec | |
| 21 | "cantor a Zero = Succ a" | |
| 22 | "cantor a (Succ b) = \<Squnion>(\<lambda>x. cantor x b) a" | |
| 23 | "cantor a (Limit f) = Limit (\<lambda>n. cantor a (f n))" | |
| 24 | ||
| 25 | consts | |
| 26 |   Nabla :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"    ("\<nabla>")
 | |
| 27 | primrec | |
| 28 | "\<nabla>f Zero = f Zero" | |
| 29 | "\<nabla>f (Succ a) = f (Succ (\<nabla>f a))" | |
| 30 | "\<nabla>f (Limit h) = Limit (\<lambda>n. \<nabla>f (h n))" | |
| 31 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 32 | definition deriv :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" where | 
| 13262 | 33 | "deriv f \<equiv> \<nabla>(\<Squnion>f)" | 
| 34 | ||
| 35 | consts | |
| 36 | veblen :: "ordinal \<Rightarrow> ordinal \<Rightarrow> ordinal" | |
| 37 | primrec | |
| 38 | "veblen Zero = \<nabla>(OpLim (power (cantor Zero)))" | |
| 39 | "veblen (Succ a) = \<nabla>(OpLim (power (veblen a)))" | |
| 40 | "veblen (Limit f) = \<nabla>(OpLim (\<lambda>n. veblen (f n)))" | |
| 41 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
16417diff
changeset | 42 | definition veb :: "ordinal \<Rightarrow> ordinal" where | 
| 13262 | 43 | "veb a \<equiv> veblen a Zero" | 
| 35437 | 44 | |
| 45 | definition epsilon0 :: ordinal    ("\<epsilon>\<^sub>0") where
 | |
| 13262 | 46 | "\<epsilon>\<^sub>0 \<equiv> veb Zero" | 
| 35437 | 47 | |
| 48 | definition Gamma0 :: ordinal    ("\<Gamma>\<^sub>0") where
 | |
| 13262 | 49 | "\<Gamma>\<^sub>0 \<equiv> Limit (\<lambda>n. (veb^n) Zero)" | 
| 50 | thm Gamma0_def | |
| 51 | ||
| 52 | end |