doc-src/TutorialI/Rules/Forward.thy
author nipkow
Fri, 02 Nov 2007 08:59:15 +0100
changeset 25261 3dc292be0b54
parent 16417 9bc16273c2d4
child 25264 7007bc8ddae4
permissions -rw-r--r--
recdef -> fun
     1 theory Forward imports Primes begin
     2 
     3 text{*\noindent
     4 Forward proof material: of, OF, THEN, simplify, rule_format.
     5 *}
     6 
     7 text{*\noindent
     8 SKIP most developments...
     9 *}
    10 
    11 (** Commutativity **)
    12 
    13 lemma is_gcd_commute: "is_gcd k m n = is_gcd k n m"
    14 apply (auto simp add: is_gcd_def);
    15 done
    16 
    17 lemma gcd_commute: "gcd m n = gcd n m"
    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
    23 
    24 lemma gcd_1 [simp]: "gcd m (Suc 0) = Suc 0"
    25 apply simp
    26 done
    27 
    28 lemma gcd_1_left [simp]: "gcd (Suc 0) m = Suc 0"
    29 apply (simp add: gcd_commute [of "Suc 0"])
    30 done
    31 
    32 text{*\noindent
    33 as far as HERE.
    34 *}
    35 
    36 text{*\noindent
    37 SKIP THIS PROOF
    38 *}
    39 
    40 lemma gcd_mult_distrib2: "k * gcd m n = gcd (k*m) (k*n)"
    41 apply (induct_tac m n rule: gcd.induct)
    42 apply (case_tac "n=0")
    43 apply simp
    44 apply (case_tac "k=0")
    45 apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
    46 done
    47 
    48 text {*
    49 @{thm[display] gcd_mult_distrib2}
    50 \rulename{gcd_mult_distrib2}
    51 *};
    52 
    53 text{*\noindent
    54 of, simplified
    55 *}
    56 
    57 
    58 lemmas gcd_mult_0 = gcd_mult_distrib2 [of k 1];
    59 lemmas gcd_mult_1 = gcd_mult_0 [simplified];
    60 
    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 
    67 text {*
    68 example using ``of'':
    69 @{thm[display] gcd_mult_distrib2 [of _ 1]}
    70 
    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 
    77 @{thm[display] gcd_mult_0}
    78 \rulename{gcd_mult_0}
    79 
    80 @{thm[display] gcd_mult_1}
    81 \rulename{gcd_mult_1}
    82 
    83 @{thm[display] sym}
    84 \rulename{sym}
    85 *};
    86 
    87 lemmas gcd_mult0 = gcd_mult_1 [THEN sym];
    88       (*not quite right: we need ?k but this gives k*)
    89 
    90 lemmas gcd_mult0' = gcd_mult_distrib2 [of k 1, simplified, THEN sym];
    91       (*better in one step!*)
    92 
    93 text {*
    94 more legible, and variables properly generalized
    95 *};
    96 
    97 lemma gcd_mult [simp]: "gcd k (k*n) = k"
    98 by (rule gcd_mult_distrib2 [of k 1, simplified, THEN sym])
    99 
   100 
   101 lemmas gcd_self0 = gcd_mult [of k 1, simplified];
   102 
   103 
   104 text {*
   105 Rules handy with THEN
   106 
   107 @{thm[display] iffD1}
   108 \rulename{iffD1}
   109 
   110 @{thm[display] iffD2}
   111 \rulename{iffD2}
   112 *};
   113 
   114 
   115 text {*
   116 again: more legible, and variables properly generalized
   117 *};
   118 
   119 lemma gcd_self [simp]: "gcd k k = k"
   120 by (rule gcd_mult [of k 1, simplified])
   121 
   122 
   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 
   133 lemma "2 \<le> u \<Longrightarrow> u*m \<noteq> Suc(u*n)"
   134 apply (intro notI)
   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 
   161 lemma relprime_dvd_mult: 
   162       "\<lbrakk> gcd k n = 1; k dvd m*n \<rbrakk> \<Longrightarrow> k dvd m"
   163 apply (insert gcd_mult_distrib2 [of m k n])
   164 apply simp
   165 apply (erule_tac t="m" in ssubst);
   166 apply simp
   167 done
   168 
   169 
   170 text {*
   171 Another example of "insert"
   172 
   173 @{thm[display] mod_div_equality}
   174 \rulename{mod_div_equality}
   175 *};
   176 
   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 *)
   180 
   181 lemma relprime_dvd_mult_iff: "gcd k n = 1 \<Longrightarrow> (k dvd m*n) = (k dvd m)";
   182 by (blast intro: relprime_dvd_mult dvd_trans)
   183 
   184 
   185 lemma relprime_20_81: "gcd 20 81 = 1";
   186 by (simp add: gcd.simps)
   187 
   188 text {*
   189 Examples of 'OF'
   190 
   191 @{thm[display] relprime_dvd_mult}
   192 \rulename{relprime_dvd_mult}
   193 
   194 @{thm[display] relprime_dvd_mult [OF relprime_20_81]}
   195 
   196 @{thm[display] dvd_refl}
   197 \rulename{dvd_refl}
   198 
   199 @{thm[display] dvd_add}
   200 \rulename{dvd_add}
   201 
   202 @{thm[display] dvd_add [OF dvd_refl dvd_refl]}
   203 
   204 @{thm[display] dvd_add [OF _ dvd_refl]}
   205 *};
   206 
   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")
   209 txt{*
   210 the tactic leaves two subgoals:
   211 @{subgoals[display,indent=0,margin=65]}
   212 *};
   213 apply blast
   214 apply (subgoal_tac "z \<noteq> 35")
   215 txt{*
   216 the tactic leaves two subgoals:
   217 @{subgoals[display,indent=0,margin=65]}
   218 *};
   219 apply arith
   220 apply force
   221 done
   222 
   223 
   224 end