doc-src/TutorialI/Rules/Forward.thy
author paulson
Mon, 12 Nov 2001 10:56:38 +0100
changeset 12156 d2758965362e
parent 11711 ecdfd237ffee
child 12390 2fa13b499975
permissions -rw-r--r--
new-style numerals without leading #, along with generic 0 and 1
     1 theory Forward = Primes:
     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 text {*
    62 @{thm[display] gcd_mult_distrib2 [of _ 1]}
    63 
    64 @{thm[display] gcd_mult_0}
    65 \rulename{gcd_mult_0}
    66 
    67 @{thm[display] gcd_mult_1}
    68 \rulename{gcd_mult_1}
    69 
    70 @{thm[display] sym}
    71 \rulename{sym}
    72 *};
    73 
    74 lemmas gcd_mult = gcd_mult_1 [THEN sym];
    75 
    76 lemmas gcd_mult = gcd_mult_distrib2 [of k 1, simplified, THEN sym];
    77       (*better in one step!*)
    78 
    79 text {*
    80 more legible
    81 *};
    82 
    83 lemma gcd_mult [simp]: "gcd(k, k*n) = k"
    84 by (rule gcd_mult_distrib2 [of k 1, simplified, THEN sym])
    85 
    86 
    87 lemmas gcd_self = gcd_mult [of k 1, simplified];
    88 
    89 
    90 text {*
    91 Rules handy with THEN
    92 
    93 @{thm[display] iffD1}
    94 \rulename{iffD1}
    95 
    96 @{thm[display] iffD2}
    97 \rulename{iffD2}
    98 *};
    99 
   100 
   101 text {*
   102 again: more legible
   103 *};
   104 
   105 lemma gcd_self [simp]: "gcd(k,k) = k"
   106 by (rule gcd_mult [of k 1, simplified])
   107 
   108 
   109 text{*
   110 NEXT SECTION: Methods for Forward Proof
   111 
   112 NEW
   113 
   114 theorem arg_cong, useful in forward steps
   115 @{thm[display] arg_cong[no_vars]}
   116 \rulename{arg_cong}
   117 *}
   118 
   119 lemma "2 \<le> u \<Longrightarrow> u*m \<noteq> Suc(u*n)"
   120 apply intro
   121 txt{*
   122 before using arg_cong
   123 @{subgoals[display,indent=0,margin=65]}
   124 *};
   125 apply (drule_tac f="\<lambda>x. x mod u" in arg_cong)
   126 txt{*
   127 after using arg_cong
   128 @{subgoals[display,indent=0,margin=65]}
   129 *};
   130 apply (simp add: mod_Suc)
   131 done
   132 
   133 text{*
   134 have just used this rule:
   135 @{thm[display] mod_Suc[no_vars]}
   136 \rulename{mod_Suc}
   137 
   138 @{thm[display] mult_le_mono1[no_vars]}
   139 \rulename{mult_le_mono1}
   140 *}
   141 
   142 
   143 text{*
   144 example of "insert"
   145 *}
   146 
   147 lemma relprime_dvd_mult: 
   148       "\<lbrakk> gcd(k,n)=1; k dvd m*n \<rbrakk> \<Longrightarrow> k dvd m"
   149 apply (insert gcd_mult_distrib2 [of m k n])
   150 apply simp
   151 apply (erule_tac t="m" in ssubst);
   152 apply simp
   153 done
   154 
   155 
   156 text {*
   157 Another example of "insert"
   158 
   159 @{thm[display] mod_div_equality}
   160 \rulename{mod_div_equality}
   161 *};
   162 
   163 (*MOVED to Force.thy, which now depends only on Divides.thy
   164 lemma div_mult_self_is_m: "0<n \<Longrightarrow> (m*n) div n = (m::nat)"
   165 *)
   166 
   167 lemma relprime_dvd_mult_iff: "gcd(k,n)=1 \<Longrightarrow> (k dvd m*n) = (k dvd m)";
   168 by (blast intro: relprime_dvd_mult dvd_trans)
   169 
   170 
   171 lemma relprime_20_81: "gcd(20,81) = 1";
   172 by (simp add: gcd.simps)
   173 
   174 text {*
   175 Examples of 'OF'
   176 
   177 @{thm[display] relprime_dvd_mult}
   178 \rulename{relprime_dvd_mult}
   179 
   180 @{thm[display] relprime_dvd_mult [OF relprime_20_81]}
   181 
   182 @{thm[display] dvd_refl}
   183 \rulename{dvd_refl}
   184 
   185 @{thm[display] dvd_add}
   186 \rulename{dvd_add}
   187 
   188 @{thm[display] dvd_add [OF dvd_refl dvd_refl]}
   189 
   190 @{thm[display] dvd_add [OF _ dvd_refl]}
   191 *};
   192 
   193 lemma "\<lbrakk>(z::int) < 37; 66 < 2*z; z*z \<noteq> 1225; Q(34); Q(36)\<rbrakk> \<Longrightarrow> Q(z)";
   194 apply (subgoal_tac "z = 34 \<or> z = 36")
   195 txt{*
   196 the tactic leaves two subgoals:
   197 @{subgoals[display,indent=0,margin=65]}
   198 *};
   199 apply blast
   200 apply (subgoal_tac "z \<noteq> 35")
   201 txt{*
   202 the tactic leaves two subgoals:
   203 @{subgoals[display,indent=0,margin=65]}
   204 *};
   205 apply arith
   206 apply force
   207 done
   208 
   209 
   210 end