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