2014-11-28 13:11:23 +00:00
|
|
|
/-
|
|
|
|
Copyright (c) 2014 Microsoft Corporation. All rights reserved.
|
|
|
|
Released under Apache 2.0 license as described in the file LICENSE.
|
|
|
|
Authors: Leonardo de Moura, Jeremy Avigad
|
|
|
|
|
|
|
|
General properties of binary operations.
|
|
|
|
-/
|
2015-04-09 18:00:59 +00:00
|
|
|
open eq.ops function
|
2014-07-25 00:46:41 +00:00
|
|
|
|
|
|
|
namespace binary
|
2014-11-28 13:11:23 +00:00
|
|
|
section
|
2014-10-09 14:13:06 +00:00
|
|
|
variable {A : Type}
|
2014-11-28 13:11:23 +00:00
|
|
|
variables (op₁ : A → A → A) (inv : A → A) (one : A)
|
|
|
|
|
2015-01-26 19:49:08 +00:00
|
|
|
local notation a * b := op₁ a b
|
|
|
|
local notation a ⁻¹ := inv a
|
2014-11-28 13:11:23 +00:00
|
|
|
|
2015-10-12 03:29:31 +00:00
|
|
|
definition commutative [reducible] := ∀a b, a * b = b * a
|
|
|
|
definition associative [reducible] := ∀a b c, (a * b) * c = a * (b * c)
|
|
|
|
definition left_identity [reducible] := ∀a, one * a = a
|
|
|
|
definition right_identity [reducible] := ∀a, a * one = a
|
|
|
|
definition left_inverse [reducible] := ∀a, a⁻¹ * a = one
|
|
|
|
definition right_inverse [reducible] := ∀a, a * a⁻¹ = one
|
|
|
|
definition left_cancelative [reducible] := ∀a b c, a * b = a * c → b = c
|
2015-04-09 17:54:28 +00:00
|
|
|
definition right_cancelative [reducible] := ∀a b c, a * b = c * b → a = c
|
|
|
|
|
|
|
|
definition inv_op_cancel_left [reducible] := ∀a b, a⁻¹ * (a * b) = b
|
|
|
|
definition op_inv_cancel_left [reducible] := ∀a b, a * (a⁻¹ * b) = b
|
|
|
|
definition inv_op_cancel_right [reducible] := ∀a b, a * b⁻¹ * b = a
|
|
|
|
definition op_inv_cancel_right [reducible] := ∀a b, a * b * b⁻¹ = a
|
2014-11-28 13:11:23 +00:00
|
|
|
|
|
|
|
variable (op₂ : A → A → A)
|
|
|
|
|
2015-01-26 19:49:08 +00:00
|
|
|
local notation a + b := op₂ a b
|
2014-11-28 13:11:23 +00:00
|
|
|
|
2015-04-09 17:54:28 +00:00
|
|
|
definition left_distributive [reducible] := ∀a b c, a * (b + c) = a * b + a * c
|
|
|
|
definition right_distributive [reducible] := ∀a b c, (a + b) * c = a * c + b * c
|
|
|
|
|
|
|
|
definition right_commutative [reducible] {B : Type} (f : B → A → B) := ∀ b a₁ a₂, f (f b a₁) a₂ = f (f b a₂) a₁
|
|
|
|
definition left_commutative [reducible] {B : Type} (f : A → B → B) := ∀ a₁ a₂ b, f a₁ (f a₂ b) = f a₂ (f a₁ b)
|
2014-09-06 18:05:07 +00:00
|
|
|
end
|
2014-07-25 00:46:41 +00:00
|
|
|
|
2015-04-22 02:13:19 +00:00
|
|
|
section
|
2014-10-09 14:13:06 +00:00
|
|
|
variable {A : Type}
|
|
|
|
variable {f : A → A → A}
|
|
|
|
variable H_comm : commutative f
|
|
|
|
variable H_assoc : associative f
|
2015-04-22 02:13:19 +00:00
|
|
|
local infixl `*` := f
|
2015-04-09 17:54:28 +00:00
|
|
|
theorem left_comm : left_commutative f :=
|
2014-09-06 18:05:07 +00:00
|
|
|
take a b c, calc
|
2014-11-28 13:11:23 +00:00
|
|
|
a*(b*c) = (a*b)*c : H_assoc
|
|
|
|
... = (b*a)*c : H_comm
|
2014-09-06 18:05:07 +00:00
|
|
|
... = b*(a*c) : H_assoc
|
2014-07-25 00:46:41 +00:00
|
|
|
|
2015-04-09 17:54:28 +00:00
|
|
|
theorem right_comm : right_commutative f :=
|
2014-09-06 18:05:07 +00:00
|
|
|
take a b c, calc
|
|
|
|
(a*b)*c = a*(b*c) : H_assoc
|
2014-11-28 13:11:23 +00:00
|
|
|
... = a*(c*b) : H_comm
|
|
|
|
... = (a*c)*b : H_assoc
|
2015-07-24 15:56:18 +00:00
|
|
|
|
|
|
|
theorem comm4 (a b c d : A) : a*b*(c*d) = a*c*(b*d) :=
|
|
|
|
calc
|
|
|
|
a*b*(c*d) = a*b*c*d : H_assoc
|
|
|
|
... = a*c*b*d : right_comm H_comm H_assoc
|
|
|
|
... = a*c*(b*d) : H_assoc
|
2014-09-06 18:05:07 +00:00
|
|
|
end
|
2014-10-09 01:45:44 +00:00
|
|
|
|
2015-04-22 02:13:19 +00:00
|
|
|
section
|
2014-10-09 14:13:06 +00:00
|
|
|
variable {A : Type}
|
|
|
|
variable {f : A → A → A}
|
|
|
|
variable H_assoc : associative f
|
2015-04-22 02:13:19 +00:00
|
|
|
local infixl `*` := f
|
2014-10-09 01:45:44 +00:00
|
|
|
theorem assoc4helper (a b c d) : (a*b)*(c*d) = a*((b*c)*d) :=
|
|
|
|
calc
|
|
|
|
(a*b)*(c*d) = a*(b*(c*d)) : H_assoc
|
2014-11-28 13:11:23 +00:00
|
|
|
... = a*((b*c)*d) : H_assoc
|
2014-10-09 01:45:44 +00:00
|
|
|
end
|
|
|
|
|
2015-04-09 18:00:59 +00:00
|
|
|
definition right_commutative_compose_right [reducible]
|
|
|
|
{A B : Type} (f : A → A → A) (g : B → A) (rcomm : right_commutative f) : right_commutative (compose_right f g) :=
|
|
|
|
λ a b₁ b₂, !rcomm
|
|
|
|
|
|
|
|
definition left_commutative_compose_left [reducible]
|
|
|
|
{A B : Type} (f : A → A → A) (g : B → A) (lcomm : left_commutative f) : left_commutative (compose_left f g) :=
|
|
|
|
λ a b₁ b₂, !lcomm
|
2014-08-07 23:59:08 +00:00
|
|
|
end binary
|