File tree Expand file tree Collapse file tree 11 files changed +15
-14
lines changed Expand file tree Collapse file tree 11 files changed +15
-14
lines changed Original file line number Diff line number Diff line change 267
267
Lemma AbsCC_square_ap_zero : forall z : CC, z [#] [0] -> AbsCC z[^]2 [#] [0].
268
268
Proof .
269
269
intros z H.
270
- astepl (Re z[^]2[+]Im z[^]2).
270
+ stepl (Re z[^]2[+]Im z[^]2).
271
271
apply (cc_inv_aid (Re z) (Im z) H).
272
272
apply AbsCC_square_Re_Im with (x := Re z) (y := Im z).
273
273
Qed .
Original file line number Diff line number Diff line change @@ -943,7 +943,7 @@ Proof.
943
943
pattern n at 1 in |- *; replace n with (1 * n).
944
944
apply degree_nexp.
945
945
apply degree_x_.
946
- replace (1 * n) with n; auto.
946
+ replace (1 * n) with n; [ auto|..] .
947
947
unfold mult in |- *.
948
948
auto with arith.
949
949
assumption.
Original file line number Diff line number Diff line change @@ -164,7 +164,7 @@ Proof.
164
164
apply less_leEq; auto.
165
165
apply a_nonneg.
166
166
apply H12.
167
- replace (n - j) with (S (n - S j)); auto with arith.
167
+ replace (n - j) with (S (n - S j)); [ auto with arith|] .
168
168
rewrite minus_Sn_m; auto with arith.
169
169
auto.
170
170
rewrite <- H22.
@@ -192,7 +192,7 @@ Proof.
192
192
intros i H15 H16.
193
193
elim (le_lt_eq_dec _ _ H15); intro H18.
194
194
apply H12.
195
- replace (n - j) with (S (n - S j)); auto with arith.
195
+ replace (n - j) with (S (n - S j)); [ auto with arith|] .
196
196
rewrite minus_Sn_m; auto with arith.
197
197
auto.
198
198
rewrite <- H18.
Original file line number Diff line number Diff line change @@ -373,7 +373,7 @@ Proof.
373
373
apply (msp_stable (msp X) (Zpos p # Pos.of_succ_nat (Init.Nat.pred (Pos.to_nat d)))). auto.
374
374
apply ball_weak_le with (1 # P_of_succ_nat z); auto.
375
375
simpl.
376
- apply Zmult_le_compat; auto with *.
376
+ apply Zmult_le_compat; [..| auto with *] .
377
377
apply Pos.le_1_l. simpl.
378
378
assert (forall i j:nat, le i j -> Pos.of_succ_nat i <= Pos.of_succ_nat j)%positive.
379
379
{ intros.
Original file line number Diff line number Diff line change @@ -682,7 +682,7 @@ Proof.
682
682
astepl (OneR[*]one_div_succ (n + (n + (n + 0)) + 2)[+]
683
683
(Two:IR)[*]one_div_succ (n + (n + (n + 0)) + 2)).
684
684
astepl ((OneR[+]Two)[*]one_div_succ (n + (n + (n + 0)) + 2)).
685
- astepl ((Three:IR)[*]one_div_succ (n + (n + (n + 0)) + 2)).
685
+ stepl ((Three:IR)[*]one_div_succ (n + (n + (n + 0)) + 2)).
686
686
2: apply mult_wdl.
687
687
2: rational.
688
688
astepr ((Three:IR)[*]([1][/] Three[//]three_ap_zero IR)[*]one_div_succ n).
Original file line number Diff line number Diff line change @@ -146,7 +146,7 @@ Proof.
146
146
apply shift_less_plus.
147
147
apply minusOne_less.
148
148
astepl (OneR[+]nexp IR p Two).
149
- astepr (OneR[+]nring (power p 2)).
149
+ stepr (OneR[+]nring (power p 2)).
150
150
apply plus_resp_eq.
151
151
apply nexp_power.
152
152
simpl in |- *.
Original file line number Diff line number Diff line change @@ -121,7 +121,7 @@ nexp IR n ([1][/] (Two:IR)[//]H).
121
121
Proof .
122
122
intros.
123
123
astepl ((zexp Two H k)[*](nexp IR (n + k) ([1][/] Two[//]H) )).
124
- astepl ((zexp Two H k)[*](zexp Two H (- (n + k)%nat))).
124
+ stepl ((zexp Two H k)[*](zexp Two H (- (n + k)%nat))).
125
125
astepr (zexp Two H (k + (- (n + k)%nat))).
126
126
apply eq_symmetric.
127
127
apply zexp_plus.
Original file line number Diff line number Diff line change @@ -567,17 +567,17 @@ Section implements_abstract_interface.
567
567
apply Qpos_nonzero.
568
568
rewrite E. simpl.
569
569
setoid_replace (wbints true #1) with (/ (1#wbints true)) by reflexivity.
570
- field... split. discriminate. apply Qpos_nonzero.
570
+ field. split. discriminate. apply Qpos_nonzero.
571
571
assert (Zpos j == (proj1_sig (ww false) / wbints false / proj1_sig x)) as jE.
572
572
{ apply (Qmult_injective_l (proj1_sig x)). apply Qpos_nonzero.
573
573
rewrite F. simpl.
574
574
setoid_replace (wbints false #1) with (/ (1#wbints false)) by reflexivity.
575
- field... split. discriminate. apply Qpos_nonzero. }
575
+ field. split. discriminate. apply Qpos_nonzero. }
576
576
assert (Zpos k == (proj1_sig totalw / w01ints / proj1_sig x)) as kE.
577
577
{ apply (Qmult_injective_l (proj1_sig x)). apply Qpos_nonzero.
578
578
rewrite G. simpl.
579
579
setoid_replace (w01ints #1) with (/ (1#w01ints)) by reflexivity.
580
- field... split. discriminate. apply Qpos_nonzero. }
580
+ field. split. discriminate. apply Qpos_nonzero. }
581
581
apply Qball_plus.
582
582
(* left case: *)
583
583
apply Σ_Qball_pos_bounds.
Original file line number Diff line number Diff line change @@ -582,7 +582,7 @@ Proof.
582
582
rewrite -> Hlr.
583
583
split; simpl.
584
584
clear - H.
585
- apply Qle_trans with 0%Q; auto with * .
585
+ apply Qle_trans with 0%Q.
586
586
apply (Qopp_le_compat 0), Qpos_nonneg.
587
587
rewrite -> Qle_minus_iff in *.
588
588
rewrite Qplus_0_r.
Original file line number Diff line number Diff line change @@ -1324,7 +1324,8 @@ Proof.
1324
1324
pose (G:=(([--][1][/]_[//]nringS_ap_zero IR n){**}([-C-][1]{-}FId){^}(S n))).
1325
1325
assert (X0:Derivative (olor [0] Two) (pos_two IR) G (([-C-][1]{-}FId){^}n)).
1326
1326
unfold G.
1327
- Derivative_Help; [|apply Derivative_scal;refine (Derivative_nth _ _ _ _ _ _);Deriv].
1327
+ eapply Derivative_wdr; simpl in |- *;
1328
+ [|apply Derivative_scal;refine (Derivative_nth _ _ _ _ _ _);Deriv].
1328
1329
FEQ.
1329
1330
repeat constructor.
1330
1331
assert (X1:Continuous (olor [0] Two) (([-C-][1]{-}FId){^}n)).
You can’t perform that action at this time.
0 commit comments