fix(library/tc_multigraph): avoid name collisions

@avigad, @fpvandoorn, @rlewis1988, @dselsam

I changed how transitive instances are named.
The motivation is to avoid a naming collision problem found by Daniel.
Before this commit, we were getting an error on the following file
tests/lean/run/collision_bug.lean.

Now, transitive instances contain the prefix "_trans_".
It makes it clear this is an internal definition and it should not be used
by users.

This change also demonstrates (again) how the `rewrite` tactic is
fragile. The problem is that the matching procedure used by it has
very little support for solving matching constraints that involving type
class instances. Eventually, we will need to reimplement `rewrite`
using the new unification procedure used in blast.

In the meantime, the workaround is to use `krewrite` (as usual).
This commit is contained in:
Leonardo de Moura 2016-02-04 13:15:42 -08:00
parent dcb35008e1
commit 42fbc63bb6
26 changed files with 64 additions and 80 deletions

View file

@ -17,7 +17,7 @@ namespace yoneda
However, we don't want to have them globally, because that will unfold the composition g ∘ f However, we don't want to have them globally, because that will unfold the composition g ∘ f
in a Category to category.category.comp g f in a Category to category.category.comp g f
-/ -/
local attribute Category.to.precategory category.to_precategory [constructor] local attribute category.to_precategory [constructor]
-- should this be defined as "yoneda_embedding Cᵒᵖ"? -- should this be defined as "yoneda_embedding Cᵒᵖ"?
definition contravariant_yoneda_embedding [constructor] [reducible] definition contravariant_yoneda_embedding [constructor] [reducible]

View file

@ -50,7 +50,7 @@ namespace category
/- yoneda preserves existing limits -/ /- yoneda preserves existing limits -/
local attribute Category.to.precategory category.to_precategory [constructor] local attribute category.to_precategory [constructor]
definition preserves_existing_limits_yoneda_embedding_lemma [constructor] (y : cone_obj F) definition preserves_existing_limits_yoneda_embedding_lemma [constructor] (y : cone_obj F)
[H : is_terminal y] {G : Cᵒᵖ ⇒ cset} (η : constant_functor I G ⟹ ɏ ∘f F) : [H : is_terminal y] {G : Cᵒᵖ ⇒ cset} (η : constant_functor I G ⟹ ɏ ∘f F) :

View file

@ -11,7 +11,6 @@ import .colimits ..constructions.set hit.set_quotient
open eq functor is_trunc sigma pi sigma.ops trunc set_quotient open eq functor is_trunc sigma pi sigma.ops trunc set_quotient
namespace category namespace category
local attribute Category.to.precategory [unfold 1]
local attribute category.to_precategory [unfold 2] local attribute category.to_precategory [unfold 2]
definition is_complete_set_cone.{u v w} [constructor] definition is_complete_set_cone.{u v w} [constructor]

View file

@ -27,7 +27,7 @@ namespace algebra
-- coercions between them anymore. -- coercions between them anymore.
-- So, an application such as (@mul A G g h) in the following definition -- So, an application such as (@mul A G g h) in the following definition
-- is type incorrect if the coercion is not added (explicitly or implicitly). -- is type incorrect if the coercion is not added (explicitly or implicitly).
local attribute group.to.has_mul [coercion] definition group.to_has_mul [coercion] {A : Type} (H : group A) : has_mul A := _
local attribute group.to_has_inv [coercion] local attribute group.to_has_inv [coercion]
universe variable l universe variable l

View file

@ -380,7 +380,8 @@ calc
... = nat_abs a : abstr_repr ... = nat_abs a : abstr_repr
theorem padd_pneg (p : × ) : padd p (pneg p) ≡ (0, 0) := theorem padd_pneg (p : × ) : padd p (pneg p) ≡ (0, 0) :=
show pr1 p + pr2 p + 0 = pr2 p + pr1 p + 0, from !nat.add_comm ▸ rfl show pr1 p + pr2 p + 0 = pr2 p + pr1 p + 0,
by rewrite [nat.add_comm (pr1 p)]
theorem padd_padd_pneg (p q : × ) : padd (padd p q) (pneg q) ≡ p := theorem padd_padd_pneg (p q : × ) : padd (padd p q) (pneg q) ≡ p :=
calc pr1 p + pr1 q + pr2 q + pr2 p calc pr1 p + pr1 q + pr2 q + pr2 p

View file

@ -71,7 +71,7 @@ theorem one_pow : ∀ n : , 1^n = (1:A)
theorem pow_add (a : A) (m n : ) : a^(m + n) = a^m * a^n := theorem pow_add (a : A) (m n : ) : a^(m + n) = a^m * a^n :=
begin begin
induction n with n ih, induction n with n ih,
{rewrite [nat.add_zero, pow_zero, mul_one]}, {krewrite [nat.add_zero, pow_zero, mul_one]},
rewrite [add_succ, *pow_succ', ih, mul.assoc] rewrite [add_succ, *pow_succ', ih, mul.assoc]
end end
@ -171,7 +171,7 @@ include s
theorem pow_pos {a : A} (H : a > 0) (n : ) : a ^ n > 0 := theorem pow_pos {a : A} (H : a > 0) (n : ) : a ^ n > 0 :=
begin begin
induction n, induction n,
rewrite pow_zero, krewrite pow_zero,
apply zero_lt_one, apply zero_lt_one,
rewrite pow_succ', rewrite pow_succ',
apply mul_pos, apply mul_pos,
@ -181,7 +181,7 @@ theorem pow_pos {a : A} (H : a > 0) (n : ) : a ^ n > 0 :=
theorem pow_ge_one_of_ge_one {a : A} (H : a ≥ 1) (n : ) : a ^ n ≥ 1 := theorem pow_ge_one_of_ge_one {a : A} (H : a ≥ 1) (n : ) : a ^ n ≥ 1 :=
begin begin
induction n, induction n,
rewrite pow_zero, krewrite pow_zero,
apply le.refl, apply le.refl,
rewrite [pow_succ', -mul_one 1], rewrite [pow_succ', -mul_one 1],
apply mul_le_mul v_0 H zero_le_one, apply mul_le_mul v_0 H zero_le_one,

View file

@ -21,7 +21,7 @@ theorem zero_pow {m : } (mpos : m > 0) : 0^m = (0 : A) :=
have h₁ : ∀ m : nat, (0 : A)^(succ m) = (0 : A), have h₁ : ∀ m : nat, (0 : A)^(succ m) = (0 : A),
begin begin
intro m, induction m, intro m, induction m,
rewrite pow_one, krewrite pow_one,
apply zero_mul apply zero_mul
end, end,
obtain m' (h₂ : m = succ m'), from exists_eq_succ_of_pos mpos, obtain m' (h₂ : m = succ m'), from exists_eq_succ_of_pos mpos,
@ -45,7 +45,7 @@ or.elim (eq_zero_or_pos m)
begin begin
intro m, intro m,
induction m with m ih, induction m with m ih,
{rewrite pow_one; intros; assumption}, {krewrite pow_one; intros; assumption},
rewrite pow_succ, rewrite pow_succ,
intro H, intro H,
cases eq_zero_or_eq_zero_of_mul_eq_zero H with h₃ h₄, cases eq_zero_or_eq_zero_of_mul_eq_zero H with h₃ h₄,
@ -73,7 +73,7 @@ or.elim (eq_zero_or_pos m)
begin begin
intro m, intro m,
induction m with m ih, induction m with m ih,
{rewrite pow_one; assumption}, { krewrite pow_one; assumption },
rewrite pow_succ, rewrite pow_succ,
apply division_ring.mul_ne_zero H ih apply division_ring.mul_ne_zero H ih
end, end,
@ -136,7 +136,7 @@ monoid_has_pow_nat
theorem abs_pow (a : A) (n : ) : abs (a^n) = abs a^n := theorem abs_pow (a : A) (n : ) : abs (a^n) = abs a^n :=
begin begin
induction n with n ih, induction n with n ih,
rewrite [*pow_zero, (abs_of_nonneg zero_le_one : abs (1 : A) = 1)], krewrite [*pow_zero, (abs_of_nonneg zero_le_one : abs (1 : A) = 1)],
rewrite [*pow_succ, abs_mul, ih] rewrite [*pow_succ, abs_mul, ih]
end end
@ -149,7 +149,7 @@ include s
theorem field.div_pow (a : A) {b : A} {n : } (bnz : b ≠ 0) : (a / b)^n = a^n / b^n := theorem field.div_pow (a : A) {b : A} {n : } (bnz : b ≠ 0) : (a / b)^n = a^n / b^n :=
begin begin
induction n with n ih, induction n with n ih,
rewrite [*pow_zero, div_one], krewrite [*pow_zero, div_one],
have bnnz : b^n ≠ 0, from division_ring.pow_ne_zero_of_ne_zero bnz, have bnnz : b^n ≠ 0, from division_ring.pow_ne_zero_of_ne_zero bnz,
rewrite [*pow_succ, ih, !field.div_mul_div bnz bnnz] rewrite [*pow_succ, ih, !field.div_mul_div bnz bnnz]
end end
@ -163,7 +163,7 @@ include s
theorem div_pow (a : A) {b : A} {n : } : (a / b)^n = a^n / b^n := theorem div_pow (a : A) {b : A} {n : } : (a / b)^n = a^n / b^n :=
begin begin
induction n with n ih, induction n with n ih,
rewrite [*pow_zero, div_one], krewrite [*pow_zero, div_one],
rewrite [*pow_succ, ih, div_mul_div] rewrite [*pow_succ, ih, div_mul_div]
end end

View file

@ -30,10 +30,10 @@ protected proposition eta (z : ) : complex.mk (complex.re z) (complex.im z) =
by cases z; exact rfl by cases z; exact rfl
definition of_real [coercion] (x : ) : := complex.mk x 0 definition of_real [coercion] (x : ) : := complex.mk x 0
definition of_rat [coercion] (q : ) : := rat.to.complex q definition of_rat [coercion] (q : ) : := q
definition of_int [coercion] (i : ) : := int.to.complex i definition of_int [coercion] (i : ) : := i
definition of_nat [coercion] (n : ) : := nat.to.complex n definition of_nat [coercion] (n : ) : := n
definition of_num [coercion] [reducible] (n : num) : := num.to.complex n definition of_num [coercion] [reducible] (n : num) : := n
protected definition prio : num := num.pred real.prio protected definition prio : num := num.pred real.prio

View file

@ -99,7 +99,7 @@ calc
... < 0 : neg_neg_of_pos this ... < 0 : neg_neg_of_pos this
protected theorem zero_div (b : ) : 0 / b = 0 := protected theorem zero_div (b : ) : 0 / b = 0 :=
by rewrite [of_nat_div_eq, nat.zero_div, of_nat_zero, mul_zero] by krewrite [of_nat_div_eq, nat.zero_div, of_nat_zero, mul_zero]
protected theorem div_zero (a : ) : a / 0 = 0 := protected theorem div_zero (a : ) : a / 0 = 0 :=
by rewrite [div_def, sign_zero, zero_mul] by rewrite [div_def, sign_zero, zero_mul]

View file

@ -23,7 +23,7 @@ theorem of_nat_pow (a n : ) : of_nat (a^n) = (of_nat a)^n :=
begin begin
induction n with n ih, induction n with n ih,
apply eq.refl, apply eq.refl,
rewrite [pow_succ, pow_succ, of_nat_mul, ih] krewrite [pow_succ, pow_succ, of_nat_mul, ih]
end end
end int end int

View file

@ -35,7 +35,7 @@ proposition sum_range_def (m n : ) (f : → A) :
proposition sum_range_self (m : ) (f : → A) : proposition sum_range_self (m : ) (f : → A) :
(∑ i = m...m, f i) = f m := (∑ i = m...m, f i) = f m :=
by rewrite [↑sum_range, succ_sub !le.refl, nat.sub_self, sum_up_to_one, zero_add] by krewrite [↑sum_range, succ_sub !le.refl, nat.sub_self, sum_up_to_one, zero_add]
proposition sum_range_succ {m n : } (f : → A) (H : m ≤ succ n) : proposition sum_range_succ {m n : } (f : → A) (H : m ≤ succ n) :
(∑ i = m...succ n, f i) = (∑ i = m...n, f i) + f (succ n) := (∑ i = m...succ n, f i) = (∑ i = m...n, f i) + f (succ n) :=
@ -126,7 +126,7 @@ proposition prod_range_def (m n : ) (f : → A) :
proposition prod_range_self (m : ) (f : → A) : proposition prod_range_self (m : ) (f : → A) :
(∏ i = m...m, f i) = f m := (∏ i = m...m, f i) = f m :=
by rewrite [↑prod_range, succ_sub !le.refl, nat.sub_self, prod_up_to_one, zero_add] by krewrite [↑prod_range, succ_sub !le.refl, nat.sub_self, prod_up_to_one, zero_add]
proposition prod_range_succ {m n : } (f : → A) (H : m ≤ succ n) : proposition prod_range_succ {m n : } (f : → A) (H : m ≤ succ n) :
(∏ i = m...succ n, f i) = (∏ i = m...n, f i) * f (succ n) := (∏ i = m...succ n, f i) = (∏ i = m...n, f i) * f (succ n) :=

View file

@ -24,7 +24,7 @@ or.elim (eq_zero_or_pos m)
begin begin
intro m, intro m,
induction m with m ih, induction m with m ih,
{rewrite pow_one; intros; assumption}, {krewrite pow_one; intros; assumption},
rewrite pow_succ, rewrite pow_succ,
intro H, intro H,
cases eq_zero_or_eq_zero_of_mul_eq_zero H with h₃ h₄, cases eq_zero_or_eq_zero_of_mul_eq_zero H with h₃ h₄,
@ -53,7 +53,7 @@ theorem le_pow_self {x : } (H : x > 1) : ∀ i, i ≤ x^i
theorem mul_self_eq_pow_2 (a : nat) : a * a = a ^ 2 := theorem mul_self_eq_pow_2 (a : nat) : a * a = a ^ 2 :=
show a * a = a ^ (succ (succ zero)), from show a * a = a ^ (succ (succ zero)), from
by rewrite [*pow_succ, *pow_zero, mul_one] by krewrite [*pow_succ, *pow_zero, mul_one]
theorem pow_cancel_left : ∀ {a b c : nat}, a > 1 → a ^ b = a ^ c → b = c theorem pow_cancel_left : ∀ {a b c : nat}, a > 1 → a ^ b = a ^ c → b = c
| a 0 0 h₁ h₂ := rfl | a 0 0 h₁ h₂ := rfl

View file

@ -353,8 +353,8 @@ namespace rat
/- operations -/ /- operations -/
definition of_int [coercion] (i : ) : := ⟦prerat.of_int i⟧ definition of_int [coercion] (i : ) : := ⟦prerat.of_int i⟧
definition of_nat [coercion] (n : ) : := nat.to.rat n definition of_nat [coercion] (n : ) : := n
definition of_num [coercion] [reducible] (n : num) : := num.to.rat n definition of_num [coercion] [reducible] (n : num) : := n
protected definition prio := num.pred int.prio protected definition prio := num.pred int.prio

View file

@ -1092,8 +1092,8 @@ has_sub.mk real.sub
open rat -- no coercions before open rat -- no coercions before
definition of_rat [coercion] (a : ) : := quot.mk (r_const a) definition of_rat [coercion] (a : ) : := quot.mk (r_const a)
definition of_int [coercion] (i : ) : := int.to.real i definition of_int [coercion] (i : ) : := i
definition of_nat [coercion] (n : ) : := nat.to.real n definition of_nat [coercion] (n : ) : := n
definition of_num [coercion] [reducible] (n : num) : := of_rat (rat.of_num n) definition of_num [coercion] [reducible] (n : num) : := of_rat (rat.of_num n)
definition real_has_zero [reducible] [instance] [priority real.prio] : has_zero real := definition real_has_zero [reducible] [instance] [priority real.prio] : has_zero real :=

View file

@ -367,11 +367,11 @@ proposition mul_converges_to_seq (HX : X ⟶ x in ) (HY : Y ⟶ y in ) :
apply converges_to_seq_constant, apply converges_to_seq_constant,
rewrite -{0}zero_add, rewrite -{0}zero_add,
apply add_converges_to_seq, apply add_converges_to_seq,
rewrite -(mul_zero K), krewrite -(mul_zero K),
apply mul_left_converges_to_seq, apply mul_left_converges_to_seq,
apply abs_sub_converges_to_seq_of_converges_to_seq, apply abs_sub_converges_to_seq_of_converges_to_seq,
exact HY, exact HY,
rewrite -(mul_zero (abs y)), krewrite -(mul_zero (abs y)),
apply mul_left_converges_to_seq, apply mul_left_converges_to_seq,
apply abs_sub_converges_to_seq_of_converges_to_seq, apply abs_sub_converges_to_seq_of_converges_to_seq,
exact HX exact HX

View file

@ -58,7 +58,7 @@ begin
induction n with [n, ih], induction n with [n, ih],
{apply rfl}, {apply rfl},
change choose (succ n) (succ 0) = succ n, change choose (succ n) (succ 0) = succ n,
rewrite [choose_succ_succ, ih, choose_zero_right] krewrite [choose_succ_succ, ih, choose_zero_right]
end end
theorem choose_pos {n : } : ∀ {k : }, k ≤ n → choose n k > 0 := theorem choose_pos {n : } : ∀ {k : }, k ≤ n → choose n k > 0 :=
@ -87,7 +87,7 @@ begin
krewrite [one_mul, choose_zero_succ, choose_eq_zero_of_lt H, zero_mul]}}, krewrite [one_mul, choose_zero_succ, choose_eq_zero_of_lt H, zero_mul]}},
intro k, intro k,
cases k with k', cases k with k',
{rewrite [choose_zero_right, choose_one_right]}, {krewrite [choose_zero_right, choose_one_right]},
rewrite [choose_succ_succ (succ n), right_distrib, -ih (succ k')], rewrite [choose_succ_succ (succ n), right_distrib, -ih (succ k')],
rewrite [choose_succ_succ at {1}, left_distrib, *succ_mul (succ n), mul_succ, -ih k'], rewrite [choose_succ_succ at {1}, left_distrib, *succ_mul (succ n), mul_succ, -ih k'],
rewrite [*add.assoc, add.left_comm (choose n _)] rewrite [*add.assoc, add.left_comm (choose n _)]

View file

@ -246,7 +246,7 @@ theorem mult_pow_eq_zero_of_prime_of_ne {p q : } (primep : prime p) (primeq :
(pneq : p ≠ q) (i : ) : mult p (q^i) = 0 := (pneq : p ≠ q) (i : ) : mult p (q^i) = 0 :=
begin begin
induction i with i ih, induction i with i ih,
{rewrite [pow_zero, mult_one_right]}, {krewrite [pow_zero, mult_one_right]},
have qpos : q > 0, from pos_of_prime primeq, have qpos : q > 0, from pos_of_prime primeq,
have qipos : q^i > 0, from !pow_pos_of_pos qpos, have qipos : q^i > 0, from !pow_pos_of_pos qpos,
rewrite [pow_succ', mult_mul primep qipos qpos, ih, mult_eq_zero_of_prime_of_ne primep rewrite [pow_succ', mult_mul primep qipos qpos, ih, mult_eq_zero_of_prime_of_ne primep

View file

@ -59,30 +59,8 @@ struct add_edge_fn {
m_graph.m_edges.insert(edge, src); m_graph.m_edges.insert(edge, src);
} }
name compose(name const & base_name, name const & e1, name const & e2) {
static name compose_name_core(name const & src, name const & tgt) { pair<environment, name> env_e = ::lean::compose(m_env, *m_tc, e2, e1, optional<name>(base_name));
return src + name("to") + tgt;
}
static name compose_name(name const & p, name const & src, name const & tgt) {
if (is_prefix_of(p, tgt))
return compose_name_core(src, tgt.replace_prefix(p, name()));
else if (p.is_atomic())
return compose_name_core(src, tgt);
else
return compose_name(p.get_prefix(), src, tgt);
}
static name compose_name(name const & src, name const & tgt) {
if (src.is_atomic())
return compose_name_core(src, tgt);
else
return compose_name(src.get_prefix(), src, tgt);
}
name compose(name const & src, name const & e1, name const & e2, name const & tgt) {
name n = compose_name(src, tgt);
pair<environment, name> env_e = ::lean::compose(m_env, *m_tc, e2, e1, optional<name>(n));
m_env = env_e.first; m_env = env_e.first;
return env_e.second; return env_e.second;
} }
@ -90,6 +68,7 @@ struct add_edge_fn {
pair<environment, list<tc_edge>> operator()(name const & src, name const & edge, name const & tgt) { pair<environment, list<tc_edge>> operator()(name const & src, name const & edge, name const & tgt) {
buffer<tc_edge> new_edges; buffer<tc_edge> new_edges;
if (auto preds = m_graph.m_predecessors.find(src)) { if (auto preds = m_graph.m_predecessors.find(src)) {
name base_name = edge.append_before("_trans_to_");
for (name const & pred : *preds) { for (name const & pred : *preds) {
if (pred == tgt) if (pred == tgt)
continue; // avoid loops continue; // avoid loops
@ -97,7 +76,7 @@ struct add_edge_fn {
for (pair<name, name> const & p : *pred_succ) { for (pair<name, name> const & p : *pred_succ) {
if (p.second != src) if (p.second != src)
continue; continue;
name new_e = compose(pred, p.first, edge, tgt); name new_e = compose(base_name, p.first, edge);
new_edges.emplace_back(pred, new_e, tgt); new_edges.emplace_back(pred, new_e, tgt);
} }
} }
@ -107,15 +86,16 @@ struct add_edge_fn {
buffer<tc_edge> new_back_edges; buffer<tc_edge> new_back_edges;
new_back_edges.append(new_edges); new_back_edges.append(new_edges);
if (auto succs = m_graph.m_successors.find(tgt)) { if (auto succs = m_graph.m_successors.find(tgt)) {
name base_name = edge.append_before("_trans_of_");
for (pair<name, name> const & p : *succs) { for (pair<name, name> const & p : *succs) {
if (src == p.second) if (src == p.second)
continue; // avoid loops continue; // avoid loops
name new_e = compose(src, edge, p.first, p.second); name new_e = compose(base_name, edge, p.first);
new_edges.emplace_back(src, new_e, p.second); new_edges.emplace_back(src, new_e, p.second);
for (auto const & back_edge : new_back_edges) { for (auto const & back_edge : new_back_edges) {
if (back_edge.m_from != p.second) if (back_edge.m_from != p.second)
continue; continue;
name new_e = compose(back_edge.m_from, back_edge.m_cnst, p.first, p.second); name new_e = compose(base_name, back_edge.m_cnst, p.first);
new_edges.emplace_back(back_edge.m_from, new_e, p.second); new_edges.emplace_back(back_edge.m_from, new_e, p.second);
} }
} }

View file

@ -2,4 +2,4 @@ import algebra.category.functor.equivalence
-- print prefix category.equivalence -- print prefix category.equivalence
check @category.equivalence.to._Fun -- check @category.equivalence.to._Fun

View file

@ -1,6 +0,0 @@
category.equivalence.to._Fun :
Π {C : category.Precategory} {D : category.Precategory} (c : category.equivalence C D)
{a b : category.Precategory.carrier C},
category.hom a b →
category.hom (functor.to_fun_ob (category.equivalence.to_functor c) a)
(functor.to_fun_ob (category.equivalence.to_functor c) b)

View file

@ -1,8 +1,15 @@
{x : ∈ S | x > 0} : set {x : ∈ S | x > 0} : set
{x : ∈ s | x > 0} : finset {x : ∈ s | x > 0} : finset
@set.sep.{1} nat (λ (x : nat), @gt.{1} nat _source.to.has_lt x (@zero.{1} nat _source.to.has_zero)) S : set.{1} nat @set.sep.{1} nat
(λ (x : nat),
@gt.{1} nat nat._trans_of_decidable_linear_ordered_semiring_13 x
(@zero.{1} nat nat._trans_of_decidable_linear_ordered_semiring_6))
S :
set.{1} nat
@finset.sep.{1} nat (λ (a b : nat), nat.has_decidable_eq a b) @finset.sep.{1} nat (λ (a b : nat), nat.has_decidable_eq a b)
(λ (x : nat), @gt.{1} nat _source.to.has_lt x (@zero.{1} nat _source.to.has_zero)) (λ (x : nat),
(λ (a : nat), nat.decidable_lt (@zero.{1} nat _source.to.has_zero) a) @gt.{1} nat nat._trans_of_decidable_linear_ordered_semiring_13 x
(@zero.{1} nat nat._trans_of_decidable_linear_ordered_semiring_6))
(λ (a : nat), nat.decidable_lt (@zero.{1} nat nat._trans_of_decidable_linear_ordered_semiring_6) a)
s : s :
finset.{1} nat finset.{1} nat

View file

@ -13,7 +13,7 @@ begin
have ceq : c = 0, begin have ceq : c = 0, begin
rewrite aeq0 at h₃, rewrite aeq0 at h₃,
rewrite add_zero at h₃, rewrite add_zero at h₃,
rewrite add_succ at h₃, krewrite add_succ at h₃,
krewrite add_zero at h₃, krewrite add_zero at h₃,
injection h₃, exact a_1 injection h₃, exact a_1
end, end,

View file

@ -1,4 +1,4 @@
@add.{1} nat _source.to.has_add a b : nat @add.{1} nat nat._trans_of_decidable_linear_ordered_semiring_2 a b : nat
@add.{1} int _source.to.has_add_1 i j : int @add.{1} int int._trans_of_linear_ordered_comm_ring_16 i j : int
@add.{1} nat _source.to.has_add a b : nat @add.{1} nat nat._trans_of_decidable_linear_ordered_semiring_2 a b : nat
@add.{1} int _source.to.has_add_1 i j : int @add.{1} int int._trans_of_linear_ordered_comm_ring_16 i j : int

View file

@ -1,8 +1,9 @@
a + of_num b = 10 : Prop a + of_num b = 10 : Prop
@eq.{1} nat @eq.{1} nat
(@add.{1} nat _source.to.has_add ((λ (x : nat), x) a) (@add.{1} nat nat._trans_of_decidable_linear_ordered_semiring_2 ((λ (x : nat), x) a)
(nat.of_num (@bit0.{1} num num_has_add (@one.{1} num num_has_one)))) (nat.of_num (@bit0.{1} num num_has_add (@one.{1} num num_has_one))))
(@bit0.{1} nat _source.to.has_add (@bit0.{1} nat nat._trans_of_decidable_linear_ordered_semiring_2
(@bit1.{1} nat _source.to.has_one _source.to.has_add (@bit1.{1} nat nat._trans_of_decidable_linear_ordered_semiring_22 nat._trans_of_decidable_linear_ordered_semiring_2
(@bit0.{1} nat _source.to.has_add (@one.{1} nat _source.to.has_one)))) : (@bit0.{1} nat nat._trans_of_decidable_linear_ordered_semiring_2
(@one.{1} nat nat._trans_of_decidable_linear_ordered_semiring_22)))) :
Prop Prop

View file

@ -0,0 +1,2 @@
import theories.measure_theory.sigma_algebra
import data.int.order

View file

@ -38,6 +38,6 @@ example (a₁ a₂ a₃ b₁ b₂ b₃ : nat) : (a₁+2, a₂+3, a₃+1) = (b₁
begin begin
intro H, injection H with a₁b₁ sa₂b₂ a₃sb₃, intro H, injection H with a₁b₁ sa₂b₂ a₃sb₃,
esimp at *, esimp at *,
rewrite [a₁b₁, a₃sb₃, sa₂b₂], krewrite [a₁b₁, a₃sb₃, -sa₂b₂],
repeat (split | esimp) repeat (split | esimp)
end end