diff -r e9ab4ad7bd15 -r 23307fd33906 src/Doc/Tutorial/Rules/Basic.thy --- a/src/Doc/Tutorial/Rules/Basic.thy Thu Jan 11 13:48:17 2018 +0100 +++ b/src/Doc/Tutorial/Rules/Basic.thy Fri Jan 12 14:08:53 2018 +0100 @@ -35,9 +35,9 @@ apply assumption done -text {* +text \ by eliminates uses of assumption and done -*} +\ lemma imp_uncurry': "P \ Q \ R \ P \ Q \ R" apply (rule impI) @@ -47,21 +47,21 @@ by (drule mp) -text {* +text \ substitution @{thm[display] ssubst} \rulename{ssubst} -*} +\ lemma "\ x = f x; P(f x) \ \ P x" by (erule ssubst) -text {* +text \ also provable by simp (re-orients) -*} +\ -text {* +text \ the subst method @{thm[display] mult.commute} @@ -69,17 +69,17 @@ this would fail: apply (simp add: mult.commute) -*} +\ lemma "\P x y z; Suc x < y\ \ f z = x*y" -txt{* +txt\ @{subgoals[display,indent=0,margin=65]} -*} +\ apply (subst mult.commute) -txt{* +txt\ @{subgoals[display,indent=0,margin=65]} -*} +\ oops (*exercise involving THEN*) @@ -90,11 +90,11 @@ lemma "\x = f x; triple (f x) (f x) x\ \ triple x x x" apply (erule ssubst) - --{* @{subgoals[display,indent=0,margin=65]} *} -back --{* @{subgoals[display,indent=0,margin=65]} *} -back --{* @{subgoals[display,indent=0,margin=65]} *} -back --{* @{subgoals[display,indent=0,margin=65]} *} -back --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ +back \\@{subgoals[display,indent=0,margin=65]}\ +back \\@{subgoals[display,indent=0,margin=65]}\ +back \\@{subgoals[display,indent=0,margin=65]}\ +back \\@{subgoals[display,indent=0,margin=65]}\ apply assumption done @@ -102,9 +102,9 @@ apply (erule ssubst, assumption) done -text{* +text\ or better still -*} +\ lemma "\ x = f x; triple (f x) (f x) x \ \ triple x x x" by (erule ssubst) @@ -120,7 +120,7 @@ by (erule_tac P="\u. triple u u x" in ssubst) -text {* +text \ negation @{thm[display] notI} @@ -143,41 +143,41 @@ @{thm[display] contrapos_nn} \rulename{contrapos_nn} -*} +\ lemma "\\(P\Q); \(R\Q)\ \ R" apply (erule_tac Q="R\Q" in contrapos_np) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (intro impI) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ by (erule notE) -text {* +text \ @{thm[display] disjCI} \rulename{disjCI} -*} +\ lemma "(P \ Q) \ R \ P \ Q \ R" apply (intro disjCI conjI) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (elim conjE disjE) apply assumption - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ by (erule contrapos_np, rule conjI) -text{* +text\ proof\ {\isacharparenleft}prove{\isacharparenright}{\isacharcolon}\ step\ {\isadigit{6}}\isanewline \isanewline goal\ {\isacharparenleft}lemma{\isacharparenright}{\isacharcolon}\isanewline {\isacharparenleft}P\ {\isasymor}\ Q{\isacharparenright}\ {\isasymand}\ R\ {\isasymLongrightarrow}\ P\ {\isasymor}\ Q\ {\isasymand}\ R\isanewline \ {\isadigit{1}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ Q\isanewline \ {\isadigit{2}}{\isachardot}\ {\isasymlbrakk}R{\isacharsemicolon}\ Q{\isacharsemicolon}\ {\isasymnot}\ P{\isasymrbrakk}\ {\isasymLongrightarrow}\ R -*} +\ -text{*rule_tac, etc.*} +text\rule_tac, etc.\ lemma "P&Q" @@ -185,23 +185,23 @@ oops -text{*unification failure trace *} +text\unification failure trace\ declare [[unify_trace_failure = true]] lemma "P(a, f(b, g(e,a), b), a) \ P(a, f(b, g(c,a), b), a)" -txt{* +txt\ @{subgoals[display,indent=0,margin=65]} apply assumption Clash: e =/= c Clash: == =/= Trueprop -*} +\ oops lemma "\x y. P(x,y) --> P(y,x)" apply auto -txt{* +txt\ @{subgoals[display,indent=0,margin=65]} apply assumption @@ -209,15 +209,15 @@ Clash: == =/= Trueprop Clash: == =/= Trueprop -*} +\ oops declare [[unify_trace_failure = false]] -text{*Quantifiers*} +text\Quantifiers\ -text {* +text \ @{thm[display] allI} \rulename{allI} @@ -226,7 +226,7 @@ @{thm[display] spec} \rulename{spec} -*} +\ lemma "\x. P x \ P x" apply (rule allI) @@ -237,74 +237,74 @@ apply (drule spec) by (drule mp) -text{*rename_tac*} +text\rename_tac\ lemma "x < y \ \x y. P x (f y)" apply (intro allI) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (rename_tac v w) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ oops lemma "\\x. P x \ P (h x); P a\ \ P(h (h a))" apply (frule spec) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (drule mp, assumption) apply (drule spec) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ by (drule mp) lemma "\\x. P x \ P (f x); P a\ \ P(f (f a))" by blast -text{* -the existential quantifier*} +text\ +the existential quantifier\ -text {* +text \ @{thm[display]"exI"} \rulename{exI} @{thm[display]"exE"} \rulename{exE} -*} +\ -text{* -instantiating quantifiers explicitly by rule_tac and erule_tac*} +text\ +instantiating quantifiers explicitly by rule_tac and erule_tac\ lemma "\\x. P x \ P (h x); P a\ \ P(h (h a))" apply (frule spec) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (drule mp, assumption) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (drule_tac x = "h a" in spec) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ by (drule mp) -text {* +text \ @{thm[display]"dvd_def"} \rulename{dvd_def} -*} +\ lemma mult_dvd_mono: "\i dvd m; j dvd n\ \ i*j dvd (m*n :: nat)" apply (simp add: dvd_def) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (erule exE) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (erule exE) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (rename_tac l) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (rule_tac x="k*l" in exI) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply simp done -text{* -Hilbert-epsilon theorems*} +text\ +Hilbert-epsilon theorems\ -text{* +text\ @{thm[display] the_equality[no_vars]} \rulename{the_equality} @@ -330,29 +330,29 @@ @{thm[display] order_antisym[no_vars]} \rulename{order_antisym} -*} +\ lemma "inv Suc (Suc n) = n" by (simp add: inv_def) -text{*but we know nothing about inv Suc 0*} +text\but we know nothing about inv Suc 0\ theorem Least_equality: "\ P (k::nat); \x. P x \ k \ x \ \ (LEAST x. P(x)) = k" apply (simp add: Least_def) -txt{* +txt\ @{subgoals[display,indent=0,margin=65]} -*} +\ apply (rule the_equality) -txt{* +txt\ @{subgoals[display,indent=0,margin=65]} first subgoal is existence; second is uniqueness -*} +\ by (auto intro: order_antisym) @@ -360,19 +360,19 @@ "(\x. \y. P x y) \ \f. \x. P x (f x)" apply (rule exI, rule allI) -txt{* +txt\ @{subgoals[display,indent=0,margin=65]} state after intro rules -*} +\ apply (drule spec, erule exE) -txt{* +txt\ @{subgoals[display,indent=0,margin=65]} applying @text{someI} automatically instantiates @{term f} to @{term "\x. SOME y. P x y"} -*} +\ by (rule someI) @@ -385,7 +385,7 @@ apply (rule exI [of _ "\x. SOME y. P x y"]) by (blast intro: someI) -text{*end of Epsilon section*} +text\end of Epsilon section\ lemma "(\x. P x) \ (\x. Q x) \ \x. P x \ Q x" @@ -433,11 +433,11 @@ lemma "\y. R y y \ \x. \y. R x y" apply (rule exI) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (rule allI) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ apply (drule spec) - --{* @{subgoals[display,indent=0,margin=65]} *} + \\@{subgoals[display,indent=0,margin=65]}\ oops lemma "\x. \y. x=y"