# HG changeset patch # User wenzelm # Date 1546530170 -3600 # Node ID 4560d1f6c4932001faf8acf80ecf9f8f1d4580ca # Parent 6f755e3cd95d709484c5628b3327412ee0162729 tuned; diff -r 6f755e3cd95d -r 4560d1f6c493 src/FOL/ex/Nat_Class.thy --- a/src/FOL/ex/Nat_Class.thy Thu Jan 03 16:42:15 2019 +0100 +++ b/src/FOL/ex/Nat_Class.thy Thu Jan 03 16:42:50 2019 +0100 @@ -7,16 +7,16 @@ begin text \ - This is an abstract version of theory \Nat\. Instead of - axiomatizing a single type \nat\ we define the class of all - these types (up to isomorphism). + This is an abstract version of theory \Nat\. Instead of axiomatizing a + single type \nat\ we define the class of all these types (up to + isomorphism). - Note: The \rec\ operator had to be made \emph{monomorphic}, - because class axioms may not contain more than one type variable. + Note: The \rec\ operator had to be made \<^emph>\monomorphic\, because class axioms + may not contain more than one type variable. \ class nat = - fixes Zero :: 'a ("0") + fixes Zero :: 'a (\0\) and Suc :: "'a \ 'a" and rec :: "'a \ 'a \ ('a \ 'a \ 'a) \ 'a" assumes induct: "P(0) \ (\x. P(x) \ P(Suc(x))) \ P(n)" @@ -26,7 +26,7 @@ and rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))" begin -definition add :: "'a \ 'a \ 'a" (infixl "+" 60) +definition add :: "'a \ 'a \ 'a" (infixl \+\ 60) where "m + n = rec(m, n, \x y. Suc(y))" lemma Suc_n_not_n: "Suc(k) \ (k::'a)"