| author | wenzelm | 
| Tue, 08 May 2007 15:36:39 +0200 | |
| changeset 22868 | c82dd66560ac | 
| parent 16417 | 9bc16273c2d4 | 
| child 25261 | 3dc292be0b54 | 
| permissions | -rw-r--r-- | 
| 16417 | 1 | theory Forward imports Primes begin | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 2 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 3 | text{*\noindent
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 4 | Forward proof material: of, OF, THEN, simplify, rule_format. | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 5 | *} | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 6 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 7 | text{*\noindent
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 8 | SKIP most developments... | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 9 | *} | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 10 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 11 | (** Commutativity **) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 12 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 13 | lemma is_gcd_commute: "is_gcd k m n = is_gcd k n m" | 
| 10958 | 14 | apply (auto simp add: is_gcd_def); | 
| 15 | done | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 16 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 17 | lemma gcd_commute: "gcd(m,n) = gcd(n,m)" | 
| 10958 | 18 | apply (rule is_gcd_unique) | 
| 19 | apply (rule is_gcd) | |
| 20 | apply (subst is_gcd_commute) | |
| 21 | apply (simp add: is_gcd) | |
| 22 | done | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 23 | |
| 11711 | 24 | lemma gcd_1 [simp]: "gcd(m, Suc 0) = Suc 0" | 
| 10958 | 25 | apply simp | 
| 26 | done | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 27 | |
| 11711 | 28 | lemma gcd_1_left [simp]: "gcd(Suc 0, m) = Suc 0" | 
| 29 | apply (simp add: gcd_commute [of "Suc 0"]) | |
| 10958 | 30 | done | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 31 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 32 | text{*\noindent
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 33 | as far as HERE. | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 34 | *} | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 35 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 36 | text{*\noindent
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 37 | SKIP THIS PROOF | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 38 | *} | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 39 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 40 | lemma gcd_mult_distrib2: "k * gcd(m,n) = gcd(k*m, k*n)" | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 41 | apply (induct_tac m n rule: gcd.induct) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 42 | apply (case_tac "n=0") | 
| 10958 | 43 | apply simp | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 44 | apply (case_tac "k=0") | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 45 | apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 46 | done | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 47 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 48 | text {*
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 49 | @{thm[display] gcd_mult_distrib2}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 50 | \rulename{gcd_mult_distrib2}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 51 | *}; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 52 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 53 | text{*\noindent
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 54 | of, simplified | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 55 | *} | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 56 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 57 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 58 | lemmas gcd_mult_0 = gcd_mult_distrib2 [of k 1]; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 59 | lemmas gcd_mult_1 = gcd_mult_0 [simplified]; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 60 | |
| 14403 | 61 | lemmas where1 = gcd_mult_distrib2 [where m=1] | 
| 62 | ||
| 63 | lemmas where2 = gcd_mult_distrib2 [where m=1 and k=1] | |
| 64 | ||
| 65 | lemmas where3 = gcd_mult_distrib2 [where m=1 and k="j+k"] | |
| 66 | ||
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 67 | text {*
 | 
| 14403 | 68 | example using ``of'': | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 69 | @{thm[display] gcd_mult_distrib2 [of _ 1]}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 70 | |
| 14403 | 71 | example using ``where'': | 
| 72 | @{thm[display] gcd_mult_distrib2 [where m=1]}
 | |
| 73 | ||
| 74 | example using ``where'', ``and'': | |
| 75 | @{thm[display] gcd_mult_distrib2 [where m=1 and k="j+k"]}
 | |
| 76 | ||
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 77 | @{thm[display] gcd_mult_0}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 78 | \rulename{gcd_mult_0}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 79 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 80 | @{thm[display] gcd_mult_1}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 81 | \rulename{gcd_mult_1}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 82 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 83 | @{thm[display] sym}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 84 | \rulename{sym}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 85 | *}; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 86 | |
| 13550 | 87 | lemmas gcd_mult0 = gcd_mult_1 [THEN sym]; | 
| 88 | (*not quite right: we need ?k but this gives k*) | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 89 | |
| 13550 | 90 | lemmas gcd_mult0' = gcd_mult_distrib2 [of k 1, simplified, THEN sym]; | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 91 | (*better in one step!*) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 92 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 93 | text {*
 | 
| 13550 | 94 | more legible, and variables properly generalized | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 95 | *}; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 96 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 97 | lemma gcd_mult [simp]: "gcd(k, k*n) = k" | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 98 | by (rule gcd_mult_distrib2 [of k 1, simplified, THEN sym]) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 99 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 100 | |
| 13550 | 101 | lemmas gcd_self0 = gcd_mult [of k 1, simplified]; | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 102 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 103 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 104 | text {*
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 105 | Rules handy with THEN | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 106 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 107 | @{thm[display] iffD1}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 108 | \rulename{iffD1}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 109 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 110 | @{thm[display] iffD2}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 111 | \rulename{iffD2}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 112 | *}; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 113 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 114 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 115 | text {*
 | 
| 13550 | 116 | again: more legible, and variables properly generalized | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 117 | *}; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 118 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 119 | lemma gcd_self [simp]: "gcd(k,k) = k" | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 120 | by (rule gcd_mult [of k 1, simplified]) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 121 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 122 | |
| 10958 | 123 | text{*
 | 
| 124 | NEXT SECTION: Methods for Forward Proof | |
| 125 | ||
| 126 | NEW | |
| 127 | ||
| 128 | theorem arg_cong, useful in forward steps | |
| 129 | @{thm[display] arg_cong[no_vars]}
 | |
| 130 | \rulename{arg_cong}
 | |
| 131 | *} | |
| 132 | ||
| 11711 | 133 | lemma "2 \<le> u \<Longrightarrow> u*m \<noteq> Suc(u*n)" | 
| 12390 | 134 | apply (intro notI) | 
| 10958 | 135 | txt{*
 | 
| 136 | before using arg_cong | |
| 137 | @{subgoals[display,indent=0,margin=65]}
 | |
| 138 | *}; | |
| 139 | apply (drule_tac f="\<lambda>x. x mod u" in arg_cong) | |
| 140 | txt{*
 | |
| 141 | after using arg_cong | |
| 142 | @{subgoals[display,indent=0,margin=65]}
 | |
| 143 | *}; | |
| 144 | apply (simp add: mod_Suc) | |
| 145 | done | |
| 146 | ||
| 147 | text{*
 | |
| 148 | have just used this rule: | |
| 149 | @{thm[display] mod_Suc[no_vars]}
 | |
| 150 | \rulename{mod_Suc}
 | |
| 151 | ||
| 152 | @{thm[display] mult_le_mono1[no_vars]}
 | |
| 153 | \rulename{mult_le_mono1}
 | |
| 154 | *} | |
| 155 | ||
| 156 | ||
| 157 | text{*
 | |
| 158 | example of "insert" | |
| 159 | *} | |
| 160 | ||
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 161 | lemma relprime_dvd_mult: | 
| 11407 | 162 | "\<lbrakk> gcd(k,n)=1; k dvd m*n \<rbrakk> \<Longrightarrow> k dvd m" | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 163 | apply (insert gcd_mult_distrib2 [of m k n]) | 
| 10958 | 164 | apply simp | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 165 | apply (erule_tac t="m" in ssubst); | 
| 10958 | 166 | apply simp | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 167 | done | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 168 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 169 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 170 | text {*
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 171 | Another example of "insert" | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 172 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 173 | @{thm[display] mod_div_equality}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 174 | \rulename{mod_div_equality}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 175 | *}; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 176 | |
| 11407 | 177 | (*MOVED to Force.thy, which now depends only on Divides.thy | 
| 178 | lemma div_mult_self_is_m: "0<n \<Longrightarrow> (m*n) div n = (m::nat)" | |
| 179 | *) | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 180 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 181 | lemma relprime_dvd_mult_iff: "gcd(k,n)=1 \<Longrightarrow> (k dvd m*n) = (k dvd m)"; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 182 | by (blast intro: relprime_dvd_mult dvd_trans) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 183 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 184 | |
| 11711 | 185 | lemma relprime_20_81: "gcd(20,81) = 1"; | 
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 186 | by (simp add: gcd.simps) | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 187 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 188 | text {*
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 189 | Examples of 'OF' | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 190 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 191 | @{thm[display] relprime_dvd_mult}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 192 | \rulename{relprime_dvd_mult}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 193 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 194 | @{thm[display] relprime_dvd_mult [OF relprime_20_81]}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 195 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 196 | @{thm[display] dvd_refl}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 197 | \rulename{dvd_refl}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 198 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 199 | @{thm[display] dvd_add}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 200 | \rulename{dvd_add}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 201 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 202 | @{thm[display] dvd_add [OF dvd_refl dvd_refl]}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 203 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 204 | @{thm[display] dvd_add [OF _ dvd_refl]}
 | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 205 | *}; | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 206 | |
| 11711 | 207 | lemma "\<lbrakk>(z::int) < 37; 66 < 2*z; z*z \<noteq> 1225; Q(34); Q(36)\<rbrakk> \<Longrightarrow> Q(z)"; | 
| 208 | apply (subgoal_tac "z = 34 \<or> z = 36") | |
| 10958 | 209 | txt{*
 | 
| 210 | the tactic leaves two subgoals: | |
| 211 | @{subgoals[display,indent=0,margin=65]}
 | |
| 212 | *}; | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 213 | apply blast | 
| 11711 | 214 | apply (subgoal_tac "z \<noteq> 35") | 
| 10958 | 215 | txt{*
 | 
| 216 | the tactic leaves two subgoals: | |
| 217 | @{subgoals[display,indent=0,margin=65]}
 | |
| 218 | *}; | |
| 10846 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 219 | apply arith | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 220 | apply force | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 221 | done | 
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 222 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 223 | |
| 
623141a08705
reformatting, and splitting the end of "Primes" to create "Forward"
 paulson parents: diff
changeset | 224 | end |