diff -r a246d390f033 -r 79d117a158bd doc-src/TutorialI/Overview/LNCS/FP1.thy --- a/doc-src/TutorialI/Overview/LNCS/FP1.thy Thu Aug 08 23:53:22 2002 +0200 +++ b/doc-src/TutorialI/Overview/LNCS/FP1.thy Fri Aug 09 11:22:18 2002 +0200 @@ -35,7 +35,6 @@ by(arith) text{*Not proved automatically because it involves multiplication:*} - lemma "n*n = n \ n=0 \ n=(1::int)" (*<*)oops(*>*) @@ -46,7 +45,6 @@ by auto - subsection{*Definitions*} consts xor :: "bool \ bool \ bool" @@ -92,13 +90,6 @@ apply (simp del: rev_rev_ident) (*<*)oops(*>*) -subsubsection{*Assumptions*} - -lemma "\ xs @ zs = ys @ xs; [] @ xs = [] @ [] \ \ ys = zs" -by simp - -lemma "\x. f x = g (f (g x)) \ f [] = f [] @ []" -by(simp (no_asm)) subsubsection{*Rewriting with Definitions*} @@ -110,10 +101,9 @@ subsubsection{*Conditional Equations*} -lemma hd_Cons_tl[simp]: "xs \ [] \ hd xs # tl xs = xs" -by(case_tac xs, simp_all) - -lemma "xs \ [] \ hd(rev xs) # tl(rev xs) = rev xs" +(*<*)thm hd_Cons_tl(*>*) +text{*A pre-proved simplification rule: @{thm hd_Cons_tl[no_vars]}*} +lemma "hd(xs @ [x]) # tl(xs @ [x]) = xs @ [x]" by simp @@ -121,24 +111,16 @@ lemma "\xs. if xs = [] then A else B" apply simp -(*<*)oops(*>*)text_raw{* \isanewline\isanewline *} - -lemma "case xs @ [] of [] \ P | y#ys \ Q ys y" -apply simp -apply(simp split: list.split) (*<*)oops(*>*) +text{*Case-expressions are only split on demand.*} subsubsection{*Arithmetic*} -text{*Is simple enough for the default arithmetic:*} +text{*Only simple arithmetic:*} lemma "\ \ m < n; m < n+(1::nat) \ \ m = n" by simp - -text{*Contains boolean combinations and needs full arithmetic:*} -lemma "\ m < n \ m < n+(1::nat) \ m = n" -apply simp -by(arith) +text{*\noindent Complex goals need @{text arith}-method.*} (*<*) subsubsection{*Tracing*}