diff -r c57ec95e7763 -r ca5029d391d1 doc-src/TutorialI/Overview/LNCS/FP1.thy --- a/doc-src/TutorialI/Overview/LNCS/FP1.thy Thu Jul 31 00:01:47 2003 +0200 +++ b/doc-src/TutorialI/Overview/LNCS/FP1.thy Thu Jul 31 14:01:04 2003 +0200 @@ -1,5 +1,13 @@ (*<*)theory FP1 = Main:(*>*) +subsection{*Quickcheck*} + +lemma "rev(xs @ ys) = rev xs @ rev ys" +quickcheck +oops + +subsection{*More Syntax*} + lemma "if xs = ys then rev xs = rev ys else rev xs \ rev ys" @@ -34,6 +42,10 @@ lemma "min i (max j k) = max (min k i) (min i (j::nat))" by(arith) +text{*Full Presburger arithmetic:*} +lemma "8 \ (n::int) \ \i j. 0\i \ 0\j \ n = 3*i + 5*j" +by(arith) + text{*Not proved automatically because it involves multiplication:*} lemma "n*n = n \ n=0 \ n=(1::int)" (*<*)oops(*>*) @@ -166,13 +178,13 @@ subsubsection{*The Compiler*} -consts comp :: "('a,'v)expr \ ('a,'v)instr list" +consts compile :: "('a,'v)expr \ ('a,'v)instr list" primrec -"comp (Cex v) = [Const v]" -"comp (Vex a) = [Load a]" -"comp (Bex f e1 e2) = (comp e2) @ (comp e1) @ [Apply f]" +"compile (Cex v) = [Const v]" +"compile (Vex a) = [Load a]" +"compile (Bex f e1 e2) = (compile e2) @ (compile e1) @ [Apply f]" -theorem "exec (comp e) s [] = [value e s]" +theorem "exec (compile e) s [] = [value e s]" (*<*)oops(*>*) @@ -228,18 +240,18 @@ subsubsection{*Nested Recursion*} -datatype tree = C "tree list" +datatype tree = Tree "tree list" text{*Some trees:*} -term "C []" -term "C [C [C []], C []]" +term "Tree []" +term "Tree [Tree [Tree []], Tree []]" consts mirror :: "tree \ tree" mirrors:: "tree list \ tree list" primrec - "mirror(C ts) = C(mirrors ts)" + "mirror(Tree ts) = Tree(mirrors ts)" "mirrors [] = []" "mirrors (t # ts) = mirrors ts @ [mirror t]" @@ -272,13 +284,48 @@ apply simp_all done +text{*The ordinals:*} +datatype ord = Zero | Succ ord | Lim "nat \ ord" + +thm ord.induct[no_vars] + +instance ord :: plus .. +instance ord :: times .. + +primrec +"a + Zero = a" +"a + Succ b = Succ(a+b)" +"a + Lim F = Lim(\n. a + F n)" + +primrec +"a * Zero = Zero" +"a * Succ b = a*b + a" +"a * Lim F = Lim(\n. a * F n)" + +text{*An example provided by Stan Wainer:*} +consts H :: "ord \ (nat \ nat) \ (nat \ nat)" +primrec +"H Zero f n = n" +"H (Succ b) f n = H b f (f n)" +"H (Lim F) f n = H (F n) f n" + +lemma [simp]: "H (a+b) f = H a f \ H b f" +apply(induct b) +apply auto +done + +lemma [simp]: "H (a*b) = H b \ H a" +apply(induct b) +apply auto +done + text{* This is \emph{not} allowed: \begin{verbatim} datatype lambda = C "lambda => lambda" \end{verbatim} \begin{exercise} -Define a datatype of ordinals and the ordinal $\Gamma_0$. +Define the ordinal $\Gamma_0$. \end{exercise} *} (*<*)end(*>*)