feat(library/data/{finset,list}/bigop.lean: generalize bigops from group to monoid

This commit is contained in:
Jeremy Avigad 2015-05-08 14:00:55 +10:00
parent 42f2fc973a
commit 6fce01385c
5 changed files with 7 additions and 9 deletions

View file

@ -1,18 +1,16 @@
/- /-
Copyright (c) 2015 Microsoft Corporation. All rights reserved. Copyright (c) 2015 Microsoft Corporation. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE. Released under Apache 2.0 license as described in the file LICENSE.
Module: data.finset
Author: Leonardo de Moura Author: Leonardo de Moura
Big operator for finite sets Big operator for finite sets.
-/ -/
import algebra.group data.finset.basic data.list.bigop import algebra.group data.finset.basic data.list.bigop
open algebra finset function binary quot subtype open algebra finset function binary quot subtype
namespace finset namespace finset
variables {A B : Type} variables {A B : Type}
variable [g : comm_group B] variable [g : comm_monoid B]
include g include g
definition bigop (s : finset A) (f : A → B) : B := definition bigop (s : finset A) (f : A → B) : B :=

View file

@ -153,7 +153,7 @@ variables {A B : Type}
definition product (s₁ : finset A) (s₂ : finset B) : finset (A × B) := definition product (s₁ : finset A) (s₂ : finset B) : finset (A × B) :=
quot.lift_on₂ s₁ s₂ quot.lift_on₂ s₁ s₂
(λ l₁ l₂, (λ l₁ l₂,
to_finset_of_nodup (list.product (elt_of l₁) (elt_of l₂)) to_finset_of_nodup (product (elt_of l₁) (elt_of l₂))
(nodup_product (has_property l₁) (has_property l₂))) (nodup_product (has_property l₁) (has_property l₂)))
(λ v₁ v₂ w₁ w₂ p₁ p₂, quot.sound (perm_product p₁ p₂)) (λ v₁ v₂ w₁ w₂ p₁ p₂, quot.sound (perm_product p₁ p₂))

View file

@ -12,7 +12,7 @@ open algebra function binary quot
namespace list namespace list
variables {A B : Type} variables {A B : Type}
variable [g : group B] variable [g : monoid B]
include g include g
definition mulf (f : A → B) : B → A → B := definition mulf (f : A → B) : B → A → B :=
@ -58,7 +58,7 @@ end list
namespace list namespace list
open perm open perm
variables {A B : Type} variables {A B : Type}
variable [g : comm_group B] variable [g : comm_monoid B]
include g include g
theorem mulf_rcomm (f : A → B) : right_commutative (mulf f) := theorem mulf_rcomm (f : A → B) : right_commutative (mulf f) :=

View file

@ -268,7 +268,7 @@ theorem zip_unzip : ∀ (l : list (A × B)), zip (pr₁ (unzip l)) (pr₂ (unzip
definition flat (l : list (list A)) : list A := definition flat (l : list (list A)) : list A :=
foldl append nil l foldl append nil l
/- cross product -/ /- product -/
section product section product
definition product : list A → list B → list (A × B) definition product : list A → list B → list (A × B)

View file

@ -130,7 +130,7 @@ end erase
section disjoint section disjoint
variable {A : Type} variable {A : Type}
definition disjoint (l₁ l₂ : list A) : Prop := ∀ a, (a ∈ l₁ → a ∈ l₂ → false) definition disjoint (l₁ l₂ : list A) : Prop := ∀ ⦃a⦄, (a ∈ l₁ → a ∈ l₂ → false)
lemma disjoint_left {l₁ l₂ : list A} : disjoint l₁ l₂ → ∀ {a}, a ∈ l₁ → a ∉ l₂ := lemma disjoint_left {l₁ l₂ : list A} : disjoint l₁ l₂ → ∀ {a}, a ∈ l₁ → a ∉ l₂ :=
λ d a, d a λ d a, d a