Skip to content

Commit 24efdbe

Browse files
authored
Merge pull request #171 from mrhaandi/auto-metas
improved auto goal selection
2 parents e13df50 + 711119c commit 24efdbe

File tree

11 files changed

+15
-14
lines changed

11 files changed

+15
-14
lines changed

complex/AbsCC.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -267,7 +267,7 @@ Qed.
267267
Lemma AbsCC_square_ap_zero : forall z : CC, z [#] [0] -> AbsCC z[^]2 [#] [0].
268268
Proof.
269269
intros z H.
270-
astepl (Re z[^]2[+]Im z[^]2).
270+
stepl (Re z[^]2[+]Im z[^]2).
271271
apply (cc_inv_aid (Re z) (Im z) H).
272272
apply AbsCC_square_Re_Im with (x := Re z) (y := Im z).
273273
Qed.

complex/NRootCC.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -943,7 +943,7 @@ Proof.
943943
pattern n at 1 in |- *; replace n with (1 * n).
944944
apply degree_nexp.
945945
apply degree_x_.
946-
replace (1 * n) with n; auto.
946+
replace (1 * n) with n; [auto|..].
947947
unfold mult in |- *.
948948
auto with arith.
949949
assumption.

fta/KeyLemma.v

Lines changed: 2 additions & 2 deletions
Original file line numberDiff line numberDiff line change
@@ -164,7 +164,7 @@ Proof.
164164
apply less_leEq; auto.
165165
apply a_nonneg.
166166
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|].
168168
rewrite minus_Sn_m; auto with arith.
169169
auto.
170170
rewrite <- H22.
@@ -192,7 +192,7 @@ Proof.
192192
intros i H15 H16.
193193
elim (le_lt_eq_dec _ _ H15); intro H18.
194194
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|].
196196
rewrite minus_Sn_m; auto with arith.
197197
auto.
198198
rewrite <- H18.

metric2/FinEnum.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -373,7 +373,7 @@ Proof.
373373
apply (msp_stable (msp X) (Zpos p # Pos.of_succ_nat (Init.Nat.pred (Pos.to_nat d)))). auto.
374374
apply ball_weak_le with (1 # P_of_succ_nat z); auto.
375375
simpl.
376-
apply Zmult_le_compat; auto with *.
376+
apply Zmult_le_compat; [..|auto with *].
377377
apply Pos.le_1_l. simpl.
378378
assert (forall i j:nat, le i j -> Pos.of_succ_nat i <= Pos.of_succ_nat j)%positive.
379379
{ intros.

metrics/CPMSTheory.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -682,7 +682,7 @@ Proof.
682682
astepl (OneR[*]one_div_succ (n + (n + (n + 0)) + 2)[+]
683683
(Two:IR)[*]one_div_succ (n + (n + (n + 0)) + 2)).
684684
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)).
686686
2: apply mult_wdl.
687687
2: rational.
688688
astepr ((Three:IR)[*]([1][/] Three[//]three_ap_zero IR)[*]one_div_succ n).

metrics/ContFunctions.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -146,7 +146,7 @@ Proof.
146146
apply shift_less_plus.
147147
apply minusOne_less.
148148
astepl (OneR[+]nexp IR p Two).
149-
astepr (OneR[+]nring (power p 2)).
149+
stepr (OneR[+]nring (power p 2)).
150150
apply plus_resp_eq.
151151
apply nexp_power.
152152
simpl in |- *.

metrics/LipExt.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -121,7 +121,7 @@ nexp IR n ([1][/] (Two:IR)[//]H).
121121
Proof.
122122
intros.
123123
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))).
125125
astepr (zexp Two H (k + (- (n + k)%nat))).
126126
apply eq_symmetric.
127127
apply zexp_plus.

ode/SimpleIntegration.v

Lines changed: 3 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -567,17 +567,17 @@ Section implements_abstract_interface.
567567
apply Qpos_nonzero.
568568
rewrite E. simpl.
569569
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.
571571
assert (Zpos j == (proj1_sig (ww false) / wbints false / proj1_sig x)) as jE.
572572
{ apply (Qmult_injective_l (proj1_sig x)). apply Qpos_nonzero.
573573
rewrite F. simpl.
574574
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. }
576576
assert (Zpos k == (proj1_sig totalw / w01ints / proj1_sig x)) as kE.
577577
{ apply (Qmult_injective_l (proj1_sig x)). apply Qpos_nonzero.
578578
rewrite G. simpl.
579579
setoid_replace (w01ints #1) with (/ (1#w01ints)) by reflexivity.
580-
field... split. discriminate. apply Qpos_nonzero. }
580+
field. split. discriminate. apply Qpos_nonzero. }
581581
apply Qball_plus.
582582
(* left case: *)
583583
apply Σ_Qball_pos_bounds.

reals/fast/Interval.v

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -582,7 +582,7 @@ Proof.
582582
rewrite -> Hlr.
583583
split; simpl.
584584
clear - H.
585-
apply Qle_trans with 0%Q; auto with *.
585+
apply Qle_trans with 0%Q.
586586
apply (Qopp_le_compat 0), Qpos_nonneg.
587587
rewrite -> Qle_minus_iff in *.
588588
rewrite Qplus_0_r.

transc/Exponential.v

Lines changed: 2 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1324,7 +1324,8 @@ Proof.
13241324
pose (G:=(([--][1][/]_[//]nringS_ap_zero IR n){**}([-C-][1]{-}FId){^}(S n))).
13251325
assert (X0:Derivative (olor [0] Two) (pos_two IR) G (([-C-][1]{-}FId){^}n)).
13261326
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].
13281329
FEQ.
13291330
repeat constructor.
13301331
assert (X1:Continuous (olor [0] Two) (([-C-][1]{-}FId){^}n)).

0 commit comments

Comments
 (0)