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