~ubuntu-branches/ubuntu/wily/coq-doc/wily

« back to all changes in this revision

Viewing changes to contrib/omega/OmegaLemmas.v

  • Committer: Bazaar Package Importer
  • Author(s): Stéphane Glondu, Stéphane Glondu, Samuel Mimram
  • Date: 2010-01-07 22:50:39 UTC
  • mfrom: (1.2.2 upstream)
  • Revision ID: james.westby@ubuntu.com-20100107225039-n3cq82589u0qt0s2
Tags: 8.2pl1-1
[ Stéphane Glondu ]
* New upstream release (Closes: #563669)
  - remove patches
* Packaging overhaul:
  - use git, advertise it in Vcs-* fields of debian/control
  - use debhelper 7 and dh with override
  - use source format 3.0 (quilt)
* debian/control:
  - set Maintainer to d-o-m, set Uploaders to Sam and myself
  - add Homepage field
  - bump Standards-Version to 3.8.3
* Register PDF documentation into doc-base
* Add debian/watch
* Update debian/copyright

[ Samuel Mimram ]
* Change coq-doc's description to mention that it provides documentation in
  pdf format, not postscript, closes: #543545.

Show diffs side-by-side

added added

removed removed

Lines of Context:
 
1
(***********************************************************************)
 
2
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team    *)
 
3
(* <O___,, *        INRIA-Rocquencourt  &  LRI-CNRS-Orsay              *)
 
4
(*   \VV/  *************************************************************)
 
5
(*    //   *      This file is distributed under the terms of the      *)
 
6
(*         *       GNU Lesser General Public License Version 2.1       *)
 
7
(***********************************************************************)
 
8
 
 
9
(*i $Id: OmegaLemmas.v 11739 2009-01-02 19:33:19Z herbelin $ i*)
 
10
 
 
11
Require Import ZArith_base.
 
12
Open Local Scope Z_scope.
 
13
 
 
14
(** Factorization lemmas *)
 
15
 
 
16
Theorem Zred_factor0 : forall n:Z, n = n * 1.
 
17
  intro x; rewrite (Zmult_1_r x); reflexivity.
 
18
Qed.
 
19
 
 
20
Theorem Zred_factor1 : forall n:Z, n + n = n * 2.
 
21
Proof.
 
22
  exact Zplus_diag_eq_mult_2.
 
23
Qed.
 
24
 
 
25
Theorem Zred_factor2 : forall n m:Z, n + n * m = n * (1 + m).
 
26
Proof.
 
27
  intros x y; pattern x at 1 in |- *; rewrite <- (Zmult_1_r x);
 
28
    rewrite <- Zmult_plus_distr_r; trivial with arith.
 
29
Qed.
 
30
 
 
31
Theorem Zred_factor3 : forall n m:Z, n * m + n = n * (1 + m).
 
32
Proof.
 
33
  intros x y; pattern x at 2 in |- *; rewrite <- (Zmult_1_r x);
 
34
    rewrite <- Zmult_plus_distr_r; rewrite Zplus_comm; 
 
35
      trivial with arith.
 
36
Qed.
 
37
 
 
38
Theorem Zred_factor4 : forall n m p:Z, n * m + n * p = n * (m + p).
 
39
Proof.
 
40
  intros x y z; symmetry  in |- *; apply Zmult_plus_distr_r.
 
41
Qed.
 
42
 
 
43
Theorem Zred_factor5 : forall n m:Z, n * 0 + m = m.
 
44
Proof.
 
45
  intros x y; rewrite <- Zmult_0_r_reverse; auto with arith.
 
46
Qed.
 
47
 
 
48
Theorem Zred_factor6 : forall n:Z, n = n + 0.
 
49
Proof.
 
50
  intro; rewrite Zplus_0_r; trivial with arith.
 
51
Qed.
 
52
 
 
53
(** Other specific variants of theorems dedicated for the Omega tactic *)
 
54
 
 
55
Lemma new_var : forall x : Z, exists y : Z, x = y.
 
56
intros x; exists x; trivial with arith. 
 
57
Qed.
 
58
 
 
59
Lemma OMEGA1 : forall x y : Z, x = y -> 0 <= x -> 0 <= y.
 
60
intros x y H; rewrite H; auto with arith.
 
61
Qed.
 
62
 
 
63
Lemma OMEGA2 : forall x y : Z, 0 <= x -> 0 <= y -> 0 <= x + y.
 
64
exact Zplus_le_0_compat.
 
65
Qed. 
 
66
 
 
67
Lemma OMEGA3 : forall x y k : Z, k > 0 -> x = y * k -> x = 0 -> y = 0.
 
68
 
 
69
intros x y k H1 H2 H3; apply (Zmult_integral_l k);
 
70
 [ unfold not in |- *; intros H4; absurd (k > 0);
 
71
    [ rewrite H4; unfold Zgt in |- *; simpl in |- *; discriminate
 
72
    | assumption ]
 
73
 | rewrite <- H2; assumption ].
 
74
Qed.
 
75
 
 
76
Lemma OMEGA4 : forall x y z : Z, x > 0 -> y > x -> z * y + x <> 0.
 
77
 
 
78
unfold not in |- *; intros x y z H1 H2 H3; cut (y > 0);
 
79
 [ intros H4; cut (0 <= z * y + x);
 
80
    [ intros H5; generalize (Zmult_le_approx y z x H4 H2 H5); intros H6;
 
81
       absurd (z * y + x > 0);
 
82
       [ rewrite H3; unfold Zgt in |- *; simpl in |- *; discriminate
 
83
       | apply Zle_gt_trans with x;
 
84
          [ pattern x at 1 in |- *; rewrite <- (Zplus_0_l x);
 
85
             apply Zplus_le_compat_r; rewrite Zmult_comm; 
 
86
             generalize H4; unfold Zgt in |- *; case y;
 
87
             [ simpl in |- *; intros H7; discriminate H7
 
88
             | intros p H7; rewrite <- (Zmult_0_r (Zpos p));
 
89
                unfold Zle in |- *; rewrite Zcompare_mult_compat; 
 
90
                exact H6
 
91
             | simpl in |- *; intros p H7; discriminate H7 ]
 
92
          | assumption ] ]
 
93
    | rewrite H3; unfold Zle in |- *; simpl in |- *; discriminate ]
 
94
 | apply Zgt_trans with x; [ assumption | assumption ] ].
 
95
Qed.
 
96
 
 
97
Lemma OMEGA5 : forall x y z : Z, x = 0 -> y = 0 -> x + y * z = 0.
 
98
 
 
99
intros x y z H1 H2; rewrite H1; rewrite H2; simpl in |- *; trivial with arith.
 
100
Qed.
 
101
 
 
102
Lemma OMEGA6 : forall x y z : Z, 0 <= x -> y = 0 -> 0 <= x + y * z.
 
103
 
 
104
intros x y z H1 H2; rewrite H2; simpl in |- *; rewrite Zplus_0_r; assumption.
 
105
Qed.
 
106
 
 
107
Lemma OMEGA7 :
 
108
 forall x y z t : Z, z > 0 -> t > 0 -> 0 <= x -> 0 <= y -> 0 <= x * z + y * t.
 
109
 
 
110
intros x y z t H1 H2 H3 H4; rewrite <- (Zplus_0_l 0); apply Zplus_le_compat;
 
111
 apply Zmult_gt_0_le_0_compat; assumption.
 
112
Qed.
 
113
 
 
114
Lemma OMEGA8 : forall x y : Z, 0 <= x -> 0 <= y -> x = - y -> x = 0.
 
115
 
 
116
intros x y H1 H2 H3; elim (Zle_lt_or_eq 0 x H1);
 
117
 [ intros H4; absurd (0 < x);
 
118
    [ change (0 >= x) in |- *; apply Zle_ge; apply Zplus_le_reg_l with y;
 
119
       rewrite H3; rewrite Zplus_opp_r; rewrite Zplus_0_r; 
 
120
       assumption
 
121
    | assumption ]
 
122
 | intros H4; rewrite H4; trivial with arith ].
 
123
Qed.
 
124
 
 
125
Lemma OMEGA9 : forall x y z t : Z, y = 0 -> x = z -> y + (- x + z) * t = 0.
 
126
 
 
127
intros x y z t H1 H2; rewrite H2; rewrite Zplus_opp_l; rewrite Zmult_0_l;
 
128
 rewrite Zplus_0_r; assumption.
 
129
Qed.
 
130
 
 
131
Lemma OMEGA10 :
 
132
 forall v c1 c2 l1 l2 k1 k2 : Z,
 
133
 (v * c1 + l1) * k1 + (v * c2 + l2) * k2 =
 
134
 v * (c1 * k1 + c2 * k2) + (l1 * k1 + l2 * k2).
 
135
 
 
136
intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
 
137
 repeat rewrite Zmult_assoc; repeat elim Zplus_assoc;
 
138
 rewrite (Zplus_permute (l1 * k1) (v * c2 * k2)); trivial with arith.
 
139
Qed.
 
140
 
 
141
Lemma OMEGA11 :
 
142
 forall v1 c1 l1 l2 k1 : Z,
 
143
 (v1 * c1 + l1) * k1 + l2 = v1 * (c1 * k1) + (l1 * k1 + l2).
 
144
 
 
145
intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
 
146
 repeat rewrite Zmult_assoc; repeat elim Zplus_assoc; 
 
147
 trivial with arith.
 
148
Qed.
 
149
 
 
150
Lemma OMEGA12 :
 
151
 forall v2 c2 l1 l2 k2 : Z,
 
152
 l1 + (v2 * c2 + l2) * k2 = v2 * (c2 * k2) + (l1 + l2 * k2).
 
153
 
 
154
intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
 
155
 repeat rewrite Zmult_assoc; repeat elim Zplus_assoc; 
 
156
 rewrite Zplus_permute; trivial with arith.
 
157
Qed.
 
158
 
 
159
Lemma OMEGA13 :
 
160
 forall (v l1 l2 : Z) (x : positive),
 
161
 v * Zpos x + l1 + (v * Zneg x + l2) = l1 + l2.
 
162
 
 
163
intros; rewrite Zplus_assoc; rewrite (Zplus_comm (v * Zpos x) l1);
 
164
 rewrite (Zplus_assoc_reverse l1); rewrite <- Zmult_plus_distr_r;
 
165
 rewrite <- Zopp_neg; rewrite (Zplus_comm (- Zneg x) (Zneg x));
 
166
 rewrite Zplus_opp_r; rewrite Zmult_0_r; rewrite Zplus_0_r;
 
167
 trivial with arith.
 
168
Qed.
 
169
 
 
170
Lemma OMEGA14 :
 
171
 forall (v l1 l2 : Z) (x : positive),
 
172
 v * Zneg x + l1 + (v * Zpos x + l2) = l1 + l2.
 
173
 
 
174
intros; rewrite Zplus_assoc; rewrite (Zplus_comm (v * Zneg x) l1);
 
175
 rewrite (Zplus_assoc_reverse l1); rewrite <- Zmult_plus_distr_r;
 
176
 rewrite <- Zopp_neg; rewrite Zplus_opp_r; rewrite Zmult_0_r;
 
177
 rewrite Zplus_0_r; trivial with arith.
 
178
Qed.
 
179
Lemma OMEGA15 :
 
180
 forall v c1 c2 l1 l2 k2 : Z,
 
181
 v * c1 + l1 + (v * c2 + l2) * k2 = v * (c1 + c2 * k2) + (l1 + l2 * k2).
 
182
 
 
183
intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
 
184
 repeat rewrite Zmult_assoc; repeat elim Zplus_assoc;
 
185
 rewrite (Zplus_permute l1 (v * c2 * k2)); trivial with arith.
 
186
Qed.
 
187
 
 
188
Lemma OMEGA16 : forall v c l k : Z, (v * c + l) * k = v * (c * k) + l * k.
 
189
 
 
190
intros; repeat rewrite Zmult_plus_distr_l || rewrite Zmult_plus_distr_r;
 
191
 repeat rewrite Zmult_assoc; repeat elim Zplus_assoc; 
 
192
 trivial with arith.
 
193
Qed.
 
194
 
 
195
Lemma OMEGA17 : forall x y z : Z, Zne x 0 -> y = 0 -> Zne (x + y * z) 0.
 
196
 
 
197
unfold Zne, not in |- *; intros x y z H1 H2 H3; apply H1;
 
198
 apply Zplus_reg_l with (y * z); rewrite Zplus_comm; 
 
199
 rewrite H3; rewrite H2; auto with arith.
 
200
Qed.
 
201
 
 
202
Lemma OMEGA18 : forall x y k : Z, x = y * k -> Zne x 0 -> Zne y 0.
 
203
 
 
204
unfold Zne, not in |- *; intros x y k H1 H2 H3; apply H2; rewrite H1;
 
205
 rewrite H3; auto with arith.
 
206
Qed.
 
207
 
 
208
Lemma OMEGA19 : forall x : Z, Zne x 0 -> 0 <= x + -1 \/ 0 <= x * -1 + -1.
 
209
 
 
210
unfold Zne in |- *; intros x H; elim (Zle_or_lt 0 x);
 
211
 [ intros H1; elim Zle_lt_or_eq with (1 := H1);
 
212
    [ intros H2; left; change (0 <= Zpred x) in |- *; apply Zsucc_le_reg;
 
213
       rewrite <- Zsucc_pred; apply Zlt_le_succ; assumption
 
214
    | intros H2; absurd (x = 0); auto with arith ]
 
215
 | intros H1; right; rewrite <- Zopp_eq_mult_neg_1; rewrite Zplus_comm;
 
216
    apply Zle_left; apply Zsucc_le_reg; simpl in |- *; 
 
217
    apply Zlt_le_succ; auto with arith ].
 
218
Qed.
 
219
 
 
220
Lemma OMEGA20 : forall x y z : Z, Zne x 0 -> y = 0 -> Zne (x + y * z) 0.
 
221
 
 
222
unfold Zne, not in |- *; intros x y z H1 H2 H3; apply H1; rewrite H2 in H3;
 
223
 simpl in H3; rewrite Zplus_0_r in H3; trivial with arith.
 
224
Qed.
 
225
 
 
226
Definition fast_Zplus_comm (x y : Z) (P : Z -> Prop)
 
227
  (H : P (y + x)) := eq_ind_r P H (Zplus_comm x y).
 
228
 
 
229
Definition fast_Zplus_assoc_reverse (n m p : Z) (P : Z -> Prop)
 
230
  (H : P (n + (m + p))) := eq_ind_r P H (Zplus_assoc_reverse n m p).
 
231
 
 
232
Definition fast_Zplus_assoc (n m p : Z) (P : Z -> Prop) 
 
233
  (H : P (n + m + p)) := eq_ind_r P H (Zplus_assoc n m p).
 
234
 
 
235
Definition fast_Zplus_permute (n m p : Z) (P : Z -> Prop)
 
236
  (H : P (m + (n + p))) := eq_ind_r P H (Zplus_permute n m p).
 
237
 
 
238
Definition fast_OMEGA10 (v c1 c2 l1 l2 k1 k2 : Z) (P : Z -> Prop)
 
239
  (H : P (v * (c1 * k1 + c2 * k2) + (l1 * k1 + l2 * k2))) :=
 
240
  eq_ind_r P H (OMEGA10 v c1 c2 l1 l2 k1 k2).
 
241
 
 
242
Definition fast_OMEGA11 (v1 c1 l1 l2 k1 : Z) (P : Z -> Prop)
 
243
  (H : P (v1 * (c1 * k1) + (l1 * k1 + l2))) :=
 
244
  eq_ind_r P H (OMEGA11 v1 c1 l1 l2 k1).
 
245
Definition fast_OMEGA12 (v2 c2 l1 l2 k2 : Z) (P : Z -> Prop)
 
246
  (H : P (v2 * (c2 * k2) + (l1 + l2 * k2))) :=
 
247
  eq_ind_r P H (OMEGA12 v2 c2 l1 l2 k2).
 
248
 
 
249
Definition fast_OMEGA15 (v c1 c2 l1 l2 k2 : Z) (P : Z -> Prop)
 
250
  (H : P (v * (c1 + c2 * k2) + (l1 + l2 * k2))) :=
 
251
  eq_ind_r P H (OMEGA15 v c1 c2 l1 l2 k2).
 
252
Definition fast_OMEGA16 (v c l k : Z) (P : Z -> Prop)
 
253
  (H : P (v * (c * k) + l * k)) := eq_ind_r P H (OMEGA16 v c l k).
 
254
 
 
255
Definition fast_OMEGA13 (v l1 l2 : Z) (x : positive) (P : Z -> Prop)
 
256
  (H : P (l1 + l2)) := eq_ind_r P H (OMEGA13 v l1 l2 x).
 
257
 
 
258
Definition fast_OMEGA14 (v l1 l2 : Z) (x : positive) (P : Z -> Prop)
 
259
  (H : P (l1 + l2)) := eq_ind_r P H (OMEGA14 v l1 l2 x).
 
260
Definition fast_Zred_factor0 (x : Z) (P : Z -> Prop) 
 
261
  (H : P (x * 1)) := eq_ind_r P H (Zred_factor0 x).
 
262
 
 
263
Definition fast_Zopp_eq_mult_neg_1 (x : Z) (P : Z -> Prop)
 
264
  (H : P (x * -1)) := eq_ind_r P H (Zopp_eq_mult_neg_1 x).
 
265
 
 
266
Definition fast_Zmult_comm (x y : Z) (P : Z -> Prop)
 
267
  (H : P (y * x)) := eq_ind_r P H (Zmult_comm x y).
 
268
 
 
269
Definition fast_Zopp_plus_distr (x y : Z) (P : Z -> Prop)
 
270
  (H : P (- x + - y)) := eq_ind_r P H (Zopp_plus_distr x y).
 
271
 
 
272
Definition fast_Zopp_involutive (x : Z) (P : Z -> Prop) (H : P x) :=
 
273
  eq_ind_r P H (Zopp_involutive x).
 
274
 
 
275
Definition fast_Zopp_mult_distr_r (x y : Z) (P : Z -> Prop) 
 
276
  (H : P (x * - y)) := eq_ind_r P H (Zopp_mult_distr_r x y).
 
277
 
 
278
Definition fast_Zmult_plus_distr_l (n m p : Z) (P : Z -> Prop)
 
279
  (H : P (n * p + m * p)) := eq_ind_r P H (Zmult_plus_distr_l n m p).
 
280
Definition fast_Zmult_opp_comm (x y : Z) (P : Z -> Prop) 
 
281
  (H : P (x * - y)) := eq_ind_r P H (Zmult_opp_comm x y).
 
282
 
 
283
Definition fast_Zmult_assoc_reverse (n m p : Z) (P : Z -> Prop)
 
284
  (H : P (n * (m * p))) := eq_ind_r P H (Zmult_assoc_reverse n m p).
 
285
 
 
286
Definition fast_Zred_factor1 (x : Z) (P : Z -> Prop) 
 
287
  (H : P (x * 2)) := eq_ind_r P H (Zred_factor1 x).
 
288
 
 
289
Definition fast_Zred_factor2 (x y : Z) (P : Z -> Prop)
 
290
  (H : P (x * (1 + y))) := eq_ind_r P H (Zred_factor2 x y).
 
291
 
 
292
Definition fast_Zred_factor3 (x y : Z) (P : Z -> Prop)
 
293
  (H : P (x * (1 + y))) := eq_ind_r P H (Zred_factor3 x y).
 
294
 
 
295
Definition fast_Zred_factor4 (x y z : Z) (P : Z -> Prop)
 
296
  (H : P (x * (y + z))) := eq_ind_r P H (Zred_factor4 x y z).
 
297
 
 
298
Definition fast_Zred_factor5 (x y : Z) (P : Z -> Prop) 
 
299
  (H : P y) := eq_ind_r P H (Zred_factor5 x y).
 
300
 
 
301
Definition fast_Zred_factor6 (x : Z) (P : Z -> Prop) 
 
302
  (H : P (x + 0)) := eq_ind_r P H (Zred_factor6 x).