refactor(library): rename strategy "msimp" ==> "inst_simp"

"inst_simp" means "instantiate simplification lemmas"
The idea is to make it clear that this strategy is *not* a simplifier.
This commit is contained in:
Leonardo de Moura 2015-12-31 12:45:48 -08:00
parent 4cf5c77575
commit b35abcc6a8
8 changed files with 60 additions and 59 deletions

View file

@ -82,24 +82,24 @@ section division_ring
theorem eq_one_div_of_mul_eq_one (H : a * b = 1) : b = 1 / a :=
assert a ≠ 0, from
suppose a = 0,
have 0 = (1:A), by msimp,
have 0 = (1:A), by inst_simp,
absurd this zero_ne_one,
by msimp
by inst_simp
theorem eq_one_div_of_mul_eq_one_left (H : b * a = 1) : b = 1 / a :=
assert a ≠ 0, from
suppose a = 0,
have 0 = (1:A), by msimp,
have 0 = (1:A), by inst_simp,
absurd this zero_ne_one,
by msimp
by inst_simp
theorem division_ring.one_div_mul_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) :
(1 / a) * (1 / b) = 1 / (b * a) :=
have (b * a) * ((1 / a) * (1 / b)) = 1, by msimp,
have (b * a) * ((1 / a) * (1 / b)) = 1, by inst_simp,
eq_one_div_of_mul_eq_one this
theorem one_div_neg_one_eq_neg_one : (1:A) / (-1) = -1 :=
have (-1) * (-1) = (1:A), by msimp,
have (-1) * (-1) = (1:A), by inst_simp,
symm (eq_one_div_of_mul_eq_one this)
theorem division_ring.one_div_neg_eq_neg_one_div (H : a ≠ 0) : 1 / (- a) = - (1 / a) :=
@ -136,7 +136,7 @@ section division_ring
theorem mul_inv_eq [simp] (Ha : a ≠ 0) (Hb : b ≠ 0) : (b * a)⁻¹ = a⁻¹ * b⁻¹ :=
eq.symm (calc
a⁻¹ * b⁻¹ = (1 / a) * (1 / b) : by msimp
a⁻¹ * b⁻¹ = (1 / a) * (1 / b) : by inst_simp
... = (1 / (b * a)) : division_ring.one_div_mul_one_div Ha Hb
... = (b * a)⁻¹ : by simp)
@ -164,7 +164,7 @@ section division_ring
theorem div_eq_one_iff_eq (a : A) {b : A} (Hb : b ≠ 0) : a / b = 1 ↔ a = b :=
iff.intro
(suppose a / b = 1, calc
a = a / b * b : by msimp
a = a / b * b : by inst_simp
... = 1 * b : this
... = b : by simp)
(suppose a = b, by simp)
@ -206,9 +206,9 @@ section field
theorem field.div_mul_right (Hb : b ≠ 0) (H : a * b ≠ 0) : a / (a * b) = 1 / b :=
assert a ≠ 0, from and.left (ne_zero_and_ne_zero_of_mul_ne_zero H),
symm (calc
1 / b = a * ((1 / a) * (1 / b)) : by msimp
1 / b = a * ((1 / a) * (1 / b)) : by inst_simp
... = a * (1 / (b * a)) : division_ring.one_div_mul_one_div this Hb
... = a * (a * b)⁻¹ : by msimp)
... = a * (a * b)⁻¹ : by inst_simp)
theorem field.div_mul_left (Ha : a ≠ 0) (H : a * b ≠ 0) : b / (a * b) = 1 / a :=
let H1 : b * a ≠ 0 := mul_ne_zero_comm H in
@ -227,7 +227,7 @@ section field
theorem field.div_mul_div (a : A) {b : A} (c : A) {d : A} (Hb : b ≠ 0) (Hd : d ≠ 0) :
(a / b) * (c / d) = (a * c) / (b * d) :=
by msimp
by inst_simp
theorem mul_div_mul_left (a : A) {b c : A} (Hb : b ≠ 0) (Hc : c ≠ 0) :
(c * a) / (c * b) = a / b :=

View file

@ -155,7 +155,7 @@ section group
by simp
theorem inv_eq_of_mul_eq_one {a b : A} (H : a * b = 1) : a⁻¹ = b :=
by msimp
by inst_simp
theorem one_inv [simp] : 1⁻¹ = (1 : A) :=
inv_eq_of_mul_eq_one (one_mul 1)
@ -164,7 +164,7 @@ section group
inv_eq_of_mul_eq_one (mul.left_inv a)
theorem inv.inj {a b : A} (H : a⁻¹ = b⁻¹) : a = b :=
by msimp
by inst_simp
theorem inv_eq_inv_iff_eq (a b : A) : a⁻¹ = b⁻¹ ↔ a = b :=
iff.intro (assume H, inv.inj H) (by simp)
@ -185,19 +185,19 @@ section group
begin apply eq_inv_of_eq_inv, symmetry, exact inv_eq_of_mul_eq_one H end
theorem mul.right_inv [simp] (a : A) : a * a⁻¹ = 1 :=
by msimp
by inst_simp
theorem mul_inv_cancel_left [simp] (a b : A) : a * (a⁻¹ * b) = b :=
by msimp
by inst_simp
theorem mul_inv_cancel_right [simp] (a b : A) : a * b * b⁻¹ = a :=
by msimp
by inst_simp
theorem mul_inv [simp] (a b : A) : (a * b)⁻¹ = b⁻¹ * a⁻¹ :=
inv_eq_of_mul_eq_one (by msimp)
inv_eq_of_mul_eq_one (by inst_simp)
theorem eq_of_mul_inv_eq_one {a b : A} (H : a * b⁻¹ = 1) : a = b :=
by msimp
by inst_simp
theorem eq_mul_inv_of_mul_eq {a b c : A} (H : a * c = b) : a = b * c⁻¹ :=
by simp
@ -230,13 +230,13 @@ section group
iff.intro eq_mul_inv_of_mul_eq mul_eq_of_eq_mul_inv
theorem mul_left_cancel {a b c : A} (H : a * b = a * c) : b = c :=
by msimp
by inst_simp
theorem mul_right_cancel {a b c : A} (H : a * b = c * b) : a = c :=
by msimp
by inst_simp
theorem mul_eq_one_of_mul_eq_one {a b : A} (H : b * a = 1) : a * b = 1 :=
by msimp
by inst_simp
theorem mul_eq_one_iff_mul_eq_one (a b : A) : a * b = 1 ↔ b * a = 1 :=
iff.intro !mul_eq_one_of_mul_eq_one !mul_eq_one_of_mul_eq_one
@ -250,19 +250,19 @@ section group
local attribute conj_by [reducible]
lemma conj_compose [simp] (f g a : A) : f ∘c g ∘c a = f*g ∘c a :=
by msimp
by inst_simp
lemma conj_id [simp] (a : A) : 1 ∘c a = a :=
by msimp
by inst_simp
lemma conj_one [simp] (g : A) : g ∘c 1 = 1 :=
by msimp
by inst_simp
lemma conj_inv_cancel [simp] (g : A) : ∀ a, g⁻¹ ∘c g ∘c a = a :=
by msimp
by inst_simp
lemma conj_inv [simp] (g : A) : ∀ a, (g ∘c a)⁻¹ = g ∘c a⁻¹ :=
by msimp
by inst_simp
lemma is_conj.refl (a : A) : a ~ a := exists.intro 1 (conj_id a)
@ -275,7 +275,7 @@ section group
assume Pab, assume Pbc,
obtain x (Px : x ∘c b = a), from Pab,
obtain y (Py : y ∘c c = b), from Pbc,
exists.intro (x*y) (by msimp)
exists.intro (x*y) (by inst_simp)
end group
@ -315,17 +315,17 @@ section add_group
by simp
theorem neg_eq_of_add_eq_zero {a b : A} (H : a + b = 0) : -a = b :=
by msimp
by inst_simp
theorem neg_zero [simp] : -0 = (0 : A) := neg_eq_of_add_eq_zero (zero_add 0)
theorem neg_neg [simp] (a : A) : -(-a) = a := neg_eq_of_add_eq_zero (add.left_inv a)
theorem eq_neg_of_add_eq_zero {a b : A} (H : a + b = 0) : a = -b :=
by msimp
by inst_simp
theorem neg.inj {a b : A} (H : -a = -b) : a = b :=
by msimp
by inst_simp
theorem neg_eq_neg_iff_eq (a b : A) : -a = -b ↔ a = b :=
iff.intro (assume H, neg.inj H) (by simp)
@ -346,10 +346,10 @@ section add_group
iff.intro !eq_neg_of_eq_neg !eq_neg_of_eq_neg
theorem add.right_inv [simp] (a : A) : a + -a = 0 :=
by msimp
by inst_simp
theorem add_neg_cancel_left [simp] (a b : A) : a + (-a + b) = b :=
by msimp
by inst_simp
theorem add_neg_cancel_right [simp] (a b : A) : a + b + -b = a :=
by simp
@ -389,10 +389,10 @@ section add_group
iff.intro eq_add_neg_of_add_eq add_eq_of_eq_add_neg
theorem add_left_cancel {a b c : A} (H : a + b = a + c) : b = c :=
by msimp
by inst_simp
theorem add_right_cancel {a b c : A} (H : a + b = c + b) : a = c :=
by msimp
by inst_simp
definition add_group.to_left_cancel_semigroup [trans_instance] [reducible] :
add_left_cancel_semigroup A :=
@ -424,7 +424,7 @@ section add_group
theorem add_sub_cancel (a b : A) : a + b - b = a := !add_neg_cancel_right
theorem eq_of_sub_eq_zero {a b : A} (H : a - b = 0) : a = b :=
by msimp
by inst_simp
theorem eq_iff_sub_eq_zero (a b : A) : a = b ↔ a - b = 0 :=
iff.intro (assume H, H ▸ !sub_self) (assume H, eq_of_sub_eq_zero H)
@ -438,12 +438,12 @@ section add_group
by simp
theorem neg_sub (a b : A) : -(a - b) = b - a :=
neg_eq_of_add_eq_zero (by msimp)
neg_eq_of_add_eq_zero (by inst_simp)
theorem add_sub (a b c : A) : a + (b - c) = a + b - c := !add.assoc⁻¹
theorem sub_add_eq_sub_sub_swap (a b c : A) : a - (b + c) = a - c - b :=
by msimp
by inst_simp
theorem sub_eq_iff_eq_add (a b c : A) : a - b = c ↔ a = c + b :=
iff.intro (assume H, eq_add_of_add_neg_eq H) (assume H, add_neg_eq_of_eq_add H)
@ -573,7 +573,7 @@ by simp
theorem bit1_add_bit1_helper [add_comm_semigroup A] [has_one A] (a b t s: A)
(H : (a + b) = t) (H2 : add1 t = s) : bit1 a + bit1 b = bit0 s :=
by msimp
by inst_simp
theorem bin_add_zero [add_monoid A] (a : A) : a + zero = a :=
by simp
@ -592,14 +592,14 @@ rfl
theorem bit1_add_one_helper [has_add A] [has_one A] (a t : A) (H : add1 (bit1 a) = t) :
bit1 a + one = t :=
by msimp
by inst_simp
theorem one_add_bit1 [add_comm_semigroup A] [has_one A] (a : A) : one + bit1 a = add1 (bit1 a) :=
by simp
theorem one_add_bit1_helper [add_comm_semigroup A] [has_one A] (a t : A)
(H : add1 (bit1 a) = t) : one + bit1 a = t :=
by msimp
by inst_simp
theorem add1_bit0 [has_add A] [has_one A] (a : A) : add1 (bit0 a) = bit1 a :=
rfl
@ -610,7 +610,7 @@ by simp
theorem add1_bit1_helper [add_comm_semigroup A] [has_one A] (a t : A) (H : add1 a = t) :
add1 (bit1 a) = bit0 t :=
by msimp
by inst_simp
theorem add1_one [has_add A] [has_one A] : add1 (one : A) = bit0 one :=
rfl

View file

@ -100,7 +100,7 @@ namespace bool
theorem band_elim_left {a b : bool} (H : a && b = tt) : a = tt :=
or.elim (dichotomy a)
(suppose a = ff,
absurd (by msimp) ff_ne_tt)
absurd (by inst_simp) ff_ne_tt)
(suppose a = tt, this)
theorem band_intro {a b : bool} (H₁ : a = tt) (H₂ : b = tt) : a && b = tt :=

View file

@ -141,14 +141,14 @@ protected theorem add_right_comm : Π (n m k : ), n + m + k = n + k + m :=
right_comm nat.add_comm nat.add_assoc
protected theorem add_left_cancel {n m k : } : n + m = n + k → m = k :=
nat.induction_on n (by simp) (by msimp)
nat.induction_on n (by simp) (by inst_simp)
protected theorem add_right_cancel {n m k : } (H : n + m = k + m) : n = k :=
have H2 : m + n = m + k, by simp,
nat.add_left_cancel H2
theorem eq_zero_of_add_eq_zero_right {n m : } : n + m = 0 → n = 0 :=
nat.induction_on n (by simp) (by msimp)
nat.induction_on n (by simp) (by inst_simp)
theorem eq_zero_of_add_eq_zero_left {n m : } (H : n + m = 0) : m = 0 :=
eq_zero_of_add_eq_zero_right (!nat.add_comm ⬝ H)

View file

@ -55,7 +55,7 @@ nat.induction_on y
(by simp)
(take y,
assume IH : (x + y * z) / z = x / z + y, calc
(x + succ y * z) / z = (x + y * z + z) / z : by msimp
(x + succ y * z) / z = (x + y * z + z) / z : by inst_simp
... = succ ((x + y * z) / z) : !add_div_self H
... = succ (x / z + y) : IH)
@ -113,17 +113,17 @@ theorem add_mod_self_left [simp] (x z : ) : (x + z) % x = z % x :=
local attribute succ_mul [simp]
theorem add_mul_mod_self [simp] (x y z : ) : (x + y * z) % z = x % z :=
nat.induction_on y (by simp) (by msimp)
nat.induction_on y (by simp) (by inst_simp)
theorem add_mul_mod_self_left [simp] (x y z : ) : (x + y * z) % y = x % y :=
by msimp
by inst_simp
theorem mul_mod_left [simp] (m n : ) : (m * n) % n = 0 :=
calc (m * n) % n = (0 + m * n) % n : by simp
... = 0 : by msimp
... = 0 : by inst_simp
theorem mul_mod_right [simp] (m n : ) : (m * n) % m = 0 :=
by msimp
by inst_simp
theorem mod_lt (x : ) {y : } (H : y > 0) : x % y < y :=
nat.case_strong_induction_on x

View file

@ -48,7 +48,7 @@ theorem succ_sub_sub_succ [simp] (n m k : ) : succ n - m - succ k = n - m - k
by simp
theorem sub_self_add [simp] (n m : ) : n - (n + m) = 0 :=
by msimp
by inst_simp
protected theorem sub.right_comm (m n k : ) : m - n - k = m - k - n :=
by simp
@ -67,13 +67,13 @@ theorem mul_pred_left [simp] (n m : ) : pred n * m = n * m - m :=
nat.induction_on n (by simp) (by simp)
theorem mul_pred_right [simp] (n m : ) : n * pred m = n * m - n :=
by msimp
by inst_simp
protected theorem mul_sub_right_distrib [simp] (n m k : ) : (n - m) * k = n * k - m * k :=
nat.induction_on m (by simp) (by simp)
protected theorem mul_sub_left_distrib [simp] (n m k : ) : n * (m - k) = n * m - n * k :=
by msimp
by inst_simp
protected theorem mul_self_sub_mul_self_eq (a b : nat) : a * a - b * b = (a + b) * (a - b) :=
by rewrite [nat.mul_sub_left_distrib, *right_distrib, mul.comm b a, add.comm (a*a) (a*b),

View file

@ -107,10 +107,11 @@ definition with_options_tac (o : expr) (t : tactic) : tactic := builtin
-- with_options_tac is just a marker for the builtin 'with_attributes' notation
definition with_attributes_tac (o : expr) (n : identifier_list) (t : tactic) : tactic := builtin
definition simp : tactic := #tactic with_options [blast.strategy "simp"] blast
definition simp_nohyps : tactic := #tactic with_options [blast.strategy "simp_nohyps"] blast
definition simp_topdown : tactic := #tactic with_options [blast.strategy "simp", simplify.top_down true] blast
definition msimp : tactic := #tactic with_options [blast.strategy "ematch_simp"] blast
definition simp : tactic := #tactic with_options [blast.strategy "simp"] blast
definition simp_nohyps : tactic := #tactic with_options [blast.strategy "simp_nohyps"] blast
definition simp_topdown : tactic := #tactic with_options [blast.strategy "simp", simplify.top_down true] blast
definition inst_simp : tactic := #tactic with_options [blast.strategy "ematch_simp"] blast
definition rec_inst_simp : tactic := #tactic with_options [blast.strategy "rec_ematch_simp"] blast
definition cases (h : expr) (ids : opt_identifier_list) : tactic := builtin

View file

@ -59,7 +59,7 @@ static optional<expr> apply_ematch_simp() {
ematch_simp_action)();
}
static optional<expr> apply_ematch_simp_rec() {
static optional<expr> apply_rec_ematch_simp() {
return rec_and_then(apply_ematch_simp)();
}
@ -110,8 +110,8 @@ optional<expr> apply_strategy() {
return apply_ematch();
} else if (s_name == "ematch_simp") {
return apply_ematch_simp();
} else if (s_name == "ematch_simp_rec") {
return apply_ematch_simp_rec();
} else if (s_name == "rec_ematch_simp") {
return apply_rec_ematch_simp();
} else if (s_name == "constructor") {
return apply_constructor();
} else if (s_name == "unit") {