diff -r e9ab4ad7bd15 -r 23307fd33906 src/Doc/Tutorial/Rules/Forward.thy --- a/src/Doc/Tutorial/Rules/Forward.thy Thu Jan 11 13:48:17 2018 +0100 +++ b/src/Doc/Tutorial/Rules/Forward.thy Fri Jan 12 14:08:53 2018 +0100 @@ -1,12 +1,12 @@ theory Forward imports TPrimes begin -text{*\noindent +text\\noindent Forward proof material: of, OF, THEN, simplify, rule_format. -*} +\ -text{*\noindent +text\\noindent SKIP most developments... -*} +\ (** Commutativity **) @@ -29,13 +29,13 @@ apply (simp add: gcd_commute [of "Suc 0"]) done -text{*\noindent +text\\noindent as far as HERE. -*} +\ -text{*\noindent +text\\noindent SKIP THIS PROOF -*} +\ lemma gcd_mult_distrib2: "k * gcd m n = gcd (k*m) (k*n)" apply (induct_tac m n rule: gcd.induct) @@ -45,14 +45,14 @@ apply simp_all done -text {* +text \ @{thm[display] gcd_mult_distrib2} \rulename{gcd_mult_distrib2} -*} +\ -text{*\noindent +text\\noindent of, simplified -*} +\ lemmas gcd_mult_0 = gcd_mult_distrib2 [of k 1] @@ -64,7 +64,7 @@ lemmas where3 = gcd_mult_distrib2 [where m=1 and k="j+k"] -text {* +text \ example using ``of'': @{thm[display] gcd_mult_distrib2 [of _ 1]} @@ -82,7 +82,7 @@ @{thm[display] sym} \rulename{sym} -*} +\ lemmas gcd_mult0 = gcd_mult_1 [THEN sym] (*not quite right: we need ?k but this gives k*) @@ -90,9 +90,9 @@ lemmas gcd_mult0' = gcd_mult_distrib2 [of k 1, simplified, THEN sym] (*better in one step!*) -text {* +text \ more legible, and variables properly generalized -*} +\ lemma gcd_mult [simp]: "gcd k (k*n) = k" by (rule gcd_mult_distrib2 [of k 1, simplified, THEN sym]) @@ -101,15 +101,15 @@ lemmas gcd_self0 = gcd_mult [of k 1, simplified] -text {* +text \ @{thm[display] gcd_mult} \rulename{gcd_mult} @{thm[display] gcd_self0} \rulename{gcd_self0} -*} +\ -text {* +text \ Rules handy with THEN @{thm[display] iffD1} @@ -117,18 +117,18 @@ @{thm[display] iffD2} \rulename{iffD2} -*} +\ -text {* +text \ again: more legible, and variables properly generalized -*} +\ lemma gcd_self [simp]: "gcd k k = k" by (rule gcd_mult [of k 1, simplified]) -text{* +text\ NEXT SECTION: Methods for Forward Proof NEW @@ -136,48 +136,48 @@ theorem arg_cong, useful in forward steps @{thm[display] arg_cong[no_vars]} \rulename{arg_cong} -*} +\ lemma "2 \ u \ u*m \ Suc(u*n)" apply (intro notI) -txt{* +txt\ before using arg_cong @{subgoals[display,indent=0,margin=65]} -*} +\ apply (drule_tac f="\x. x mod u" in arg_cong) -txt{* +txt\ after using arg_cong @{subgoals[display,indent=0,margin=65]} -*} +\ apply (simp add: mod_Suc) done -text{* +text\ have just used this rule: @{thm[display] mod_Suc[no_vars]} \rulename{mod_Suc} @{thm[display] mult_le_mono1[no_vars]} \rulename{mult_le_mono1} -*} +\ -text{* +text\ example of "insert" -*} +\ lemma relprime_dvd_mult: "\ gcd k n = 1; k dvd m*n \ \ k dvd m" apply (insert gcd_mult_distrib2 [of m k n]) -txt{*@{subgoals[display,indent=0,margin=65]}*} +txt\@{subgoals[display,indent=0,margin=65]}\ apply simp -txt{*@{subgoals[display,indent=0,margin=65]}*} +txt\@{subgoals[display,indent=0,margin=65]}\ apply (erule_tac t="m" in ssubst) apply simp done -text {* +text \ @{thm[display] relprime_dvd_mult} \rulename{relprime_dvd_mult} @@ -185,7 +185,7 @@ @{thm[display] div_mult_mod_eq} \rulename{div_mult_mod_eq} -*} +\ (*MOVED to Force.thy, which now depends only on Divides.thy lemma div_mult_self_is_m: "0 (m*n) div n = (m::nat)" @@ -197,7 +197,7 @@ lemma relprime_20_81: "gcd 20 81 = 1" by (simp add: gcd.simps) -text {* +text \ Examples of 'OF' @{thm[display] relprime_dvd_mult} @@ -214,20 +214,20 @@ @{thm[display] dvd_add [OF dvd_refl dvd_refl]} @{thm[display] dvd_add [OF _ dvd_refl]} -*} +\ lemma "\(z::int) < 37; 66 < 2*z; z*z \ 1225; Q(34); Q(36)\ \ Q(z)" apply (subgoal_tac "z = 34 \ z = 36") -txt{* +txt\ the tactic leaves two subgoals: @{subgoals[display,indent=0,margin=65]} -*} +\ apply blast apply (subgoal_tac "z \ 35") -txt{* +txt\ the tactic leaves two subgoals: @{subgoals[display,indent=0,margin=65]} -*} +\ apply arith apply force done