diff -r 65d18005d802 -r 9e888d60d3e5 doc-src/TutorialI/Ifexpr/Ifexpr.thy --- a/doc-src/TutorialI/Ifexpr/Ifexpr.thy Fri Jan 05 18:32:33 2001 +0100 +++ b/doc-src/TutorialI/Ifexpr/Ifexpr.thy Fri Jan 05 18:32:57 2001 +0100 @@ -36,12 +36,12 @@ values: *} -consts value :: "boolex \\ (nat \\ bool) \\ bool"; +consts value :: "boolex \ (nat \ bool) \ bool"; primrec "value (Const b) env = b" "value (Var x) env = env x" -"value (Neg b) env = (\\ value b env)" -"value (And b c) env = (value b env \\ value c env)"; +"value (Neg b) env = (\ value b env)" +"value (And b c) env = (value b env \ value c env)"; text{*\noindent \subsubsection{If-expressions} @@ -58,7 +58,7 @@ The evaluation if If-expressions proceeds as for @{typ"boolex"}: *} -consts valif :: "ifex \\ (nat \\ bool) \\ bool"; +consts valif :: "ifex \ (nat \ bool) \ bool"; primrec "valif (CIF b) env = b" "valif (VIF x) env = env x" @@ -73,7 +73,7 @@ translate from @{typ"boolex"} into @{typ"ifex"}: *} -consts bool2if :: "boolex \\ ifex"; +consts bool2if :: "boolex \ ifex"; primrec "bool2if (Const b) = CIF b" "bool2if (Var x) = VIF x" @@ -107,13 +107,13 @@ primitive recursive functions perform this task: *} -consts normif :: "ifex \\ ifex \\ ifex \\ ifex"; +consts normif :: "ifex \ ifex \ ifex \ ifex"; primrec "normif (CIF b) t e = IF (CIF b) t e" "normif (VIF x) t e = IF (VIF x) t e" "normif (IF b t e) u f = normif b (normif t u f) (normif e u f)"; -consts norm :: "ifex \\ ifex"; +consts norm :: "ifex \ ifex"; primrec "norm (CIF b) = CIF b" "norm (VIF x) = VIF x" @@ -133,7 +133,7 @@ *} lemma [simp]: - "\\t e. valif (normif b t e) env = valif (IF b t e) env"; + "\t e. valif (normif b t e) env = valif (IF b t e) env"; (*<*) apply(induct_tac b); by(auto); @@ -150,19 +150,19 @@ the above sense? We define a function that tests If-expressions for normality *} -consts normal :: "ifex \\ bool"; +consts normal :: "ifex \ bool"; primrec "normal(CIF b) = True" "normal(VIF x) = True" -"normal(IF b t e) = (normal t \\ normal e \\ - (case b of CIF b \\ True | VIF x \\ True | IF x y z \\ False))"; +"normal(IF b t e) = (normal t \ normal e \ + (case b of CIF b \ True | VIF x \ True | IF x y z \ False))"; text{*\noindent and prove @{term"normal(norm b)"}. Of course, this requires a lemma about normality of @{term"normif"}: *} -lemma[simp]: "\\t e. normal(normif b t e) = (normal t \\ normal e)"; +lemma[simp]: "\t e. normal(normif b t e) = (normal t \ normal e)"; (*<*) apply(induct_tac b); by(auto); @@ -173,9 +173,9 @@ (*>*) text{*\medskip -How does one come up with the required lemmas? Try to prove the main theorems -without them and study carefully what @{text auto} leaves unproved. This has -to provide the clue. The necessity of universal quantification +How do we come up with the required lemmas? Try to prove the main theorems +without them and study carefully what @{text auto} leaves unproved. This +can provide the clue. The necessity of universal quantification (@{text"\t e"}) in the two lemmas is explained in \S\ref{sec:InductionHeuristics}