doc-src/TutorialI/Rules/Forward.thy
author paulson
Wed, 08 Aug 2001 14:50:28 +0200
changeset 11480 0fba0357c04c
parent 11407 138919f1a135
child 11711 ecdfd237ffee
permissions -rw-r--r--
Getting it working again with 1' instead of 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,1') = 1'"
    25 apply simp
    26 done
    27 
    28 lemma gcd_1_left [simp]: "gcd(1',m) = 1'"
    29 apply (simp add: gcd_commute [of "1'"])
    30 done
    31 
    32 text{*\noindent
    33 as far as HERE.
    34 *}
    35 
    36 
    37 text {*
    38 @{thm[display] gcd_1}
    39 \rulename{gcd_1}
    40 
    41 @{thm[display] gcd_1_left}
    42 \rulename{gcd_1_left}
    43 *};
    44 
    45 text{*\noindent
    46 SKIP THIS PROOF
    47 *}
    48 
    49 lemma gcd_mult_distrib2: "k * gcd(m,n) = gcd(k*m, k*n)"
    50 apply (induct_tac m n rule: gcd.induct)
    51 apply (case_tac "n=0")
    52 apply simp
    53 apply (case_tac "k=0")
    54 apply (simp_all add: mod_geq gcd_non_0 mod_mult_distrib2)
    55 done
    56 
    57 text {*
    58 @{thm[display] gcd_mult_distrib2}
    59 \rulename{gcd_mult_distrib2}
    60 *};
    61 
    62 text{*\noindent
    63 of, simplified
    64 *}
    65 
    66 
    67 lemmas gcd_mult_0 = gcd_mult_distrib2 [of k 1];
    68 lemmas gcd_mult_1 = gcd_mult_0 [simplified];
    69 
    70 text {*
    71 @{thm[display] gcd_mult_distrib2 [of _ 1]}
    72 
    73 @{thm[display] gcd_mult_0}
    74 \rulename{gcd_mult_0}
    75 
    76 @{thm[display] gcd_mult_1}
    77 \rulename{gcd_mult_1}
    78 
    79 @{thm[display] sym}
    80 \rulename{sym}
    81 *};
    82 
    83 lemmas gcd_mult = gcd_mult_1 [THEN sym];
    84 
    85 lemmas gcd_mult = gcd_mult_distrib2 [of k 1, simplified, THEN sym];
    86       (*better in one step!*)
    87 
    88 text {*
    89 more legible
    90 *};
    91 
    92 lemma gcd_mult [simp]: "gcd(k, k*n) = k"
    93 by (rule gcd_mult_distrib2 [of k 1, simplified, THEN sym])
    94 
    95 
    96 lemmas gcd_self = gcd_mult [of k 1, simplified];
    97 
    98 
    99 text {*
   100 Rules handy with THEN
   101 
   102 @{thm[display] iffD1}
   103 \rulename{iffD1}
   104 
   105 @{thm[display] iffD2}
   106 \rulename{iffD2}
   107 *};
   108 
   109 
   110 text {*
   111 again: more legible
   112 *};
   113 
   114 lemma gcd_self [simp]: "gcd(k,k) = k"
   115 by (rule gcd_mult [of k 1, simplified])
   116 
   117 
   118 text{*
   119 NEXT SECTION: Methods for Forward Proof
   120 
   121 NEW
   122 
   123 theorem arg_cong, useful in forward steps
   124 @{thm[display] arg_cong[no_vars]}
   125 \rulename{arg_cong}
   126 *}
   127 
   128 lemma "#2 \<le> u \<Longrightarrow> u*m \<noteq> Suc(u*n)"
   129 apply intro
   130 txt{*
   131 before using arg_cong
   132 @{subgoals[display,indent=0,margin=65]}
   133 *};
   134 apply (drule_tac f="\<lambda>x. x mod u" in arg_cong)
   135 txt{*
   136 after using arg_cong
   137 @{subgoals[display,indent=0,margin=65]}
   138 *};
   139 apply (simp add: mod_Suc)
   140 done
   141 
   142 text{*
   143 have just used this rule:
   144 @{thm[display] mod_Suc[no_vars]}
   145 \rulename{mod_Suc}
   146 
   147 @{thm[display] mult_le_mono1[no_vars]}
   148 \rulename{mult_le_mono1}
   149 *}
   150 
   151 
   152 text{*
   153 example of "insert"
   154 *}
   155 
   156 lemma relprime_dvd_mult: 
   157       "\<lbrakk> gcd(k,n)=1; k dvd m*n \<rbrakk> \<Longrightarrow> k dvd m"
   158 apply (insert gcd_mult_distrib2 [of m k n])
   159 apply simp
   160 apply (erule_tac t="m" in ssubst);
   161 apply simp
   162 done
   163 
   164 
   165 text {*
   166 Another example of "insert"
   167 
   168 @{thm[display] mod_div_equality}
   169 \rulename{mod_div_equality}
   170 *};
   171 
   172 (*MOVED to Force.thy, which now depends only on Divides.thy
   173 lemma div_mult_self_is_m: "0<n \<Longrightarrow> (m*n) div n = (m::nat)"
   174 *)
   175 
   176 lemma relprime_dvd_mult_iff: "gcd(k,n)=1 \<Longrightarrow> (k dvd m*n) = (k dvd m)";
   177 by (blast intro: relprime_dvd_mult dvd_trans)
   178 
   179 
   180 lemma relprime_20_81: "gcd(#20,#81) = 1";
   181 by (simp add: gcd.simps)
   182 
   183 text {*
   184 Examples of 'OF'
   185 
   186 @{thm[display] relprime_dvd_mult}
   187 \rulename{relprime_dvd_mult}
   188 
   189 @{thm[display] relprime_dvd_mult [OF relprime_20_81]}
   190 
   191 @{thm[display] dvd_refl}
   192 \rulename{dvd_refl}
   193 
   194 @{thm[display] dvd_add}
   195 \rulename{dvd_add}
   196 
   197 @{thm[display] dvd_add [OF dvd_refl dvd_refl]}
   198 
   199 @{thm[display] dvd_add [OF _ dvd_refl]}
   200 *};
   201 
   202 lemma "\<lbrakk>(z::int) < #37; #66 < #2*z; z*z \<noteq> #1225; Q(#34); Q(#36)\<rbrakk> \<Longrightarrow> Q(z)";
   203 apply (subgoal_tac "z = #34 \<or> z = #36")
   204 txt{*
   205 the tactic leaves two subgoals:
   206 @{subgoals[display,indent=0,margin=65]}
   207 *};
   208 apply blast
   209 apply (subgoal_tac "z \<noteq> #35")
   210 txt{*
   211 the tactic leaves two subgoals:
   212 @{subgoals[display,indent=0,margin=65]}
   213 *};
   214 apply arith
   215 apply force
   216 done
   217 
   218 
   219 end