refactor(library/algebra/group_bigops.lean,library/data/nat/bigops.lean): simplify naming scheme for bigops

This commit is contained in:
Jeremy Avigad 2015-05-17 15:24:37 +10:00
parent 6dc1cfca3c
commit 4764f6e8ec
2 changed files with 185 additions and 213 deletions

View file

@ -3,43 +3,32 @@ Copyright (c) 2015 Jeremy Avigad. 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.
Authors: Leonardo de Moura, Jeremy Avigad Authors: Leonardo de Moura, Jeremy Avigad
Finite products on a monoid, and finite sums on an additive monoid. These are called Finite products on a monoid, and finite sums on an additive monoid.
algebra.list.prod
algebra.finset.prod
algebra.list.sum
algebra.finset.sum
So when we open algebra we have list.prod etc., and when we migrate to nat, we have
nat.list.prod etc.
We have to be careful with dependencies. This theory imports files from finset and list, which We have to be careful with dependencies. This theory imports files from finset and list, which
import basic files from nat. Then nat imports this file to instantiate finite products and sums. import basic files from nat. Then nat imports this file to instantiate finite products and sums.
-/ -/
import .group data.list.basic data.list.perm data.finset.basic import .group data.list.basic data.list.perm data.finset.basic
open algebra function binary quot subtype open algebra function binary quot subtype list finset
namespace algebra namespace algebra
variables {A B : Type}
variable [deceqA : decidable_eq A]
/- list.prod -/ /- Prodl: product indexed by a list -/
namespace list -- i.e. algebra.list section monoid
open list -- i.e. ordinary lists variable [mB : monoid B]
variable {A : Type}
section monoid
variables {B : Type} [mB : monoid B]
include mB include mB
definition mulf (f : A → B) : B → A → B := definition mulf (f : A → B) : B → A → B :=
λ b a, b * f a λ b a, b * f a
definition prod (l : list A) (f : A → B) : B := definition Prodl (l : list A) (f : A → B) : B :=
list.foldl (mulf f) 1 l list.foldl (mulf f) 1 l
-- ∏ x ← l, f x -- ∏ x ← l, f x
notation `∏` binders `←` l, r:(scoped f, prod l f) := r notation `∏` binders `←` l, r:(scoped f, Prodl l f) := r
private theorem foldl_const (f : A → B) : private theorem foldl_const (f : A → B) :
∀ (l : list A) (b : B), foldl (mulf f) b l = b * foldl (mulf f) 1 l ∀ (l : list A) (b : B), foldl (mulf f) b l = b * foldl (mulf f) 1 l
@ -47,179 +36,159 @@ namespace list -- i.e. algebra.list
| (a::l) b := by rewrite [*foldl_cons, foldl_const, {foldl _ (mulf f 1 a) _}foldl_const, ↑mulf, | (a::l) b := by rewrite [*foldl_cons, foldl_const, {foldl _ (mulf f 1 a) _}foldl_const, ↑mulf,
one_mul, mul.assoc] one_mul, mul.assoc]
theorem prod_nil (f : A → B) : prod [] f = 1 := rfl theorem Prodl_nil (f : A → B) : Prodl [] f = 1 := rfl
theorem prod_cons (f : A → B) (a : A) (l : list A) : prod (a::l) f = f a * prod l f := theorem Prodl_cons (f : A → B) (a : A) (l : list A) : Prodl (a::l) f = f a * Prodl l f :=
by rewrite [↑prod, foldl_cons, foldl_const, ↑mulf, one_mul] by rewrite [↑Prodl, foldl_cons, foldl_const, ↑mulf, one_mul]
theorem prod_append : theorem Prodl_append :
∀ (l₁ l₂ : list A) (f : A → B), prod (l₁++l₂) f = prod l₁ f * prod l₂ f ∀ (l₁ l₂ : list A) (f : A → B), Prodl (l₁++l₂) f = Prodl l₁ f * Prodl l₂ f
| [] l₂ f := by rewrite [append_nil_left, prod_nil, one_mul] | [] l₂ f := by rewrite [append_nil_left, Prodl_nil, one_mul]
| (a::l) l₂ f := by rewrite [append_cons, *prod_cons, prod_append, mul.assoc] | (a::l) l₂ f := by rewrite [append_cons, *Prodl_cons, Prodl_append, mul.assoc]
section decidable_eq section deceqA
variable [H : decidable_eq A] include deceqA
include H
theorem prod_insert_of_mem (f : A → B) {a : A} {l : list A} : a ∈ l → theorem Prodl_insert_of_mem (f : A → B) {a : A} {l : list A} : a ∈ l →
prod (insert a l) f = prod l f := Prodl (insert a l) f = Prodl l f :=
assume ainl, by rewrite [insert_eq_of_mem ainl] assume ainl, by rewrite [insert_eq_of_mem ainl]
theorem prod_insert_of_not_mem (f : A → B) {a : A} {l : list A} : theorem Prodl_insert_of_not_mem (f : A → B) {a : A} {l : list A} :
a ∉ l → prod (insert a l) f = f a * prod l f := a ∉ l → Prodl (insert a l) f = f a * Prodl l f :=
assume nainl, by rewrite [insert_eq_of_not_mem nainl, prod_cons] assume nainl, by rewrite [insert_eq_of_not_mem nainl, Prodl_cons]
theorem prod_union {l₁ l₂ : list A} (f : A → B) (d : disjoint l₁ l₂) : theorem Prodl_union {l₁ l₂ : list A} (f : A → B) (d : disjoint l₁ l₂) :
prod (union l₁ l₂) f = prod l₁ f * prod l₂ f := Prodl (union l₁ l₂) f = Prodl l₁ f * Prodl l₂ f :=
by rewrite [union_eq_append d, prod_append] by rewrite [union_eq_append d, Prodl_append]
end decidable_eq end deceqA
end monoid end monoid
section comm_monoid section comm_monoid
variables {B : Type} [cmB : comm_monoid B]
include cmB
theorem prod_mul (l : list A) (f g : A → B) : prod l (λx, f x * g x) = prod l f * prod l g :=
list.induction_on l
(by rewrite [*prod_nil, mul_one])
(take a l,
assume IH,
by rewrite [*prod_cons, IH, *mul.assoc, mul.left_comm (prod l f)])
end comm_monoid
end list
/- finset.prod -/
namespace finset
open finset
variables {A B : Type}
variable [cmB : comm_monoid B] variable [cmB : comm_monoid B]
include cmB include cmB
theorem mulf_rcomm (f : A → B) : right_commutative (list.mulf f) := theorem Prodl_mul (l : list A) (f g : A → B) : Prodl l (λx, f x * g x) = Prodl l f * Prodl l g :=
list.induction_on l
(by rewrite [*Prodl_nil, mul_one])
(take a l,
assume IH,
by rewrite [*Prodl_cons, IH, *mul.assoc, mul.left_comm (Prodl l f)])
end comm_monoid
/- Prod: product indexed by a finset -/
section comm_monoid
variable [cmB : comm_monoid B]
include cmB
theorem mulf_rcomm (f : A → B) : right_commutative (mulf f) :=
right_commutative_compose_right (@has_mul.mul B cmB) f (@mul.right_comm B cmB) right_commutative_compose_right (@has_mul.mul B cmB) f (@mul.right_comm B cmB)
section theorem Prodl_eq_Prodl_of_perm (f : A → B) {l₁ l₂ : list A} :
open perm perm l₁ l₂ → Prodl l₁ f = Prodl l₂ f :=
theorem listprod_of_perm (f : A → B) {l₁ l₂ : list A} : λ p, perm.foldl_eq_of_perm (mulf_rcomm f) p 1
l₁ ~ l₂ → list.prod l₁ f = list.prod l₂ f :=
λ p, foldl_eq_of_perm (mulf_rcomm f) p 1
end
definition prod (s : finset A) (f : A → B) : B := definition Prod (s : finset A) (f : A → B) : B :=
quot.lift_on s quot.lift_on s
(λ l, list.prod (elt_of l) f) (λ l, Prodl (elt_of l) f)
(λ l₁ l₂ p, listprod_of_perm f p) (λ l₁ l₂ p, Prodl_eq_Prodl_of_perm f p)
-- ∏ x ∈ s, f x -- ∏ x ∈ s, f x
notation `∏` binders `∈` s, r:(scoped f, prod s f) := r notation `∏` binders `∈` s, r:(scoped f, prod s f) := r
theorem prod_empty (f : A → B) : prod ∅ f = 1 := theorem Prod_empty (f : A → B) : Prod ∅ f = 1 :=
list.prod_nil f Prodl_nil f
section decidable_eq section decidable_eq
variable [H : decidable_eq A] variable [H : decidable_eq A]
include H include H
theorem prod_insert_of_mem (f : A → B) {a : A} {s : finset A} : theorem Prod_insert_of_mem (f : A → B) {a : A} {s : finset A} :
a ∈ s → prod (insert a s) f = prod s f := a ∈ s → Prod (insert a s) f = Prod s f :=
quot.induction_on s quot.induction_on s
(λ l ainl, list.prod_insert_of_mem f ainl) (λ l ainl, Prodl_insert_of_mem f ainl)
theorem prod_insert_of_not_mem (f : A → B) {a : A} {s : finset A} : theorem Prod_insert_of_not_mem (f : A → B) {a : A} {s : finset A} :
a ∉ s → prod (insert a s) f = f a * prod s f := a ∉ s → Prod (insert a s) f = f a * Prod s f :=
quot.induction_on s quot.induction_on s
(λ l nainl, list.prod_insert_of_not_mem f nainl) (λ l nainl, Prodl_insert_of_not_mem f nainl)
theorem prod_union (f : A → B) {s₁ s₂ : finset A} (disj : s₁ ∩ s₂ = ∅) : theorem Prod_union (f : A → B) {s₁ s₂ : finset A} (disj : s₁ ∩ s₂ = ∅) :
prod (s₁ s₂) f = prod s₁ f * prod s₂ f := Prod (s₁ s₂) f = Prod s₁ f * Prod s₂ f :=
have H1 : disjoint s₁ s₂ → prod (s₁ s₂) f = prod s₁ f * prod s₂ f, from have H1 : disjoint s₁ s₂ → Prod (s₁ s₂) f = Prod s₁ f * Prod s₂ f, from
quot.induction_on₂ s₁ s₂ quot.induction_on₂ s₁ s₂
(λ l₁ l₂ d, list.prod_union f d), (λ l₁ l₂ d, Prodl_union f d),
H1 (disjoint_of_inter_empty disj) H1 (disjoint_of_inter_empty disj)
end decidable_eq end decidable_eq
theorem prod_mul (s : finset A) (f g : A → B) : prod s (λx, f x * g x) = prod s f * prod s g := theorem Prod_mul (s : finset A) (f g : A → B) : Prod s (λx, f x * g x) = Prod s f * Prod s g :=
quot.induction_on s (take u, !list.prod_mul) quot.induction_on s (take u, !Prodl_mul)
end comm_monoid
end finset section add_monoid
variable [amB : add_monoid B]
/- list.sum -/
namespace list
open list
variable {A : Type}
section add_monoid
variables {B : Type} [amB : add_monoid B]
include amB include amB
local attribute add_monoid.to_monoid [instance] local attribute add_monoid.to_monoid [instance]
definition sum (l : list A) (f : A → B) : B := prod l f definition Suml (l : list A) (f : A → B) : B := Prodl l f
-- ∑ x ← l, f x -- ∑ x ← l, f x
notation `∑` binders `←` l, r:(scoped f, sum l f) := r notation `∑` binders `←` l, r:(scoped f, Suml l f) := r
theorem sum_nil (f : A → B) : sum [] f = 0 := prod_nil f theorem Suml_nil (f : A → B) : Suml [] f = 0 := Prodl_nil f
theorem sum_cons (f : A → B) (a : A) (l : list A) : sum (a::l) f = f a + sum l f := theorem Suml_cons (f : A → B) (a : A) (l : list A) : Suml (a::l) f = f a + Suml l f :=
prod_cons f a l Prodl_cons f a l
theorem sum_append (l₁ l₂ : list A) (f : A → B) : sum (l₁++l₂) f = sum l₁ f + sum l₂ f := theorem Suml_append (l₁ l₂ : list A) (f : A → B) : Suml (l₁++l₂) f = Suml l₁ f + Suml l₂ f :=
prod_append l₁ l₂ f Prodl_append l₁ l₂ f
section decidable_eq section decidable_eq
variable [H : decidable_eq A] variable [H : decidable_eq A]
include H include H
theorem sum_insert_of_mem (f : A → B) {a : A} {l : list A} (H : a ∈ l) : theorem Suml_insert_of_mem (f : A → B) {a : A} {l : list A} (H : a ∈ l) :
sum (insert a l) f = sum l f := prod_insert_of_mem f H Suml (insert a l) f = Suml l f := Prodl_insert_of_mem f H
theorem sum_insert_of_not_mem (f : A → B) {a : A} {l : list A} (H : a ∉ l) : theorem Suml_insert_of_not_mem (f : A → B) {a : A} {l : list A} (H : a ∉ l) :
sum (insert a l) f = f a + sum l f := prod_insert_of_not_mem f H Suml (insert a l) f = f a + Suml l f := Prodl_insert_of_not_mem f H
theorem sum_union {l₁ l₂ : list A} (f : A → B) (d : disjoint l₁ l₂) : theorem Suml_union {l₁ l₂ : list A} (f : A → B) (d : disjoint l₁ l₂) :
sum (union l₁ l₂) f = sum l₁ f + sum l₂ f := prod_union f d Suml (union l₁ l₂) f = Suml l₁ f + Suml l₂ f := Prodl_union f d
end decidable_eq end decidable_eq
end add_monoid end add_monoid
section add_comm_monoid section add_comm_monoid
variables {B : Type} [acmB : add_comm_monoid B] variable [acmB : add_comm_monoid B]
include acmB include acmB
local attribute add_comm_monoid.to_comm_monoid [instance] local attribute add_comm_monoid.to_comm_monoid [instance]
theorem sum_add (l : list A) (f g : A → B) : sum l (λx, f x + g x) = sum l f + sum l g := theorem Suml_add (l : list A) (f g : A → B) : Suml l (λx, f x + g x) = Suml l f + Suml l g :=
prod_mul l f g Prodl_mul l f g
end add_comm_monoid end add_comm_monoid
end list
/- finset.sum -/ /- Sum -/
namespace finset section add_comm_monoid
open finset variable [acmB : add_comm_monoid B]
variable {A : Type}
section add_comm_monoid
variables {B : Type} [acmB : add_comm_monoid B]
include acmB include acmB
local attribute add_comm_monoid.to_comm_monoid [instance] local attribute add_comm_monoid.to_comm_monoid [instance]
definition sum (s : finset A) (f : A → B) : B := prod s f definition Sum (s : finset A) (f : A → B) : B := Prod s f
-- ∑ x ∈ s, f x -- ∑ x ∈ s, f x
notation `∑` binders `∈` s, r:(scoped f, sum s f) := r notation `∑` binders `∈` s, r:(scoped f, Sum s f) := r
theorem sum_empty (f : A → B) : sum ∅ f = 0 := prod_empty f theorem Sum_empty (f : A → B) : Sum ∅ f = 0 := Prod_empty f
section decidable_eq section decidable_eq
variable [H : decidable_eq A] variable [H : decidable_eq A]
include H include H
theorem sum_insert_of_mem (f : A → B) {a : A} {s : finset A} (H : a ∈ s) : theorem Sum_insert_of_mem (f : A → B) {a : A} {s : finset A} (H : a ∈ s) :
sum (insert a s) f = sum s f := prod_insert_of_mem f H Sum (insert a s) f = Sum s f := Prod_insert_of_mem f H
theorem sum_insert_of_not_mem (f : A → B) {a : A} {s : finset A} (H : a ∉ s) : theorem Sum_insert_of_not_mem (f : A → B) {a : A} {s : finset A} (H : a ∉ s) :
sum (insert a s) f = f a + sum s f := prod_insert_of_not_mem f H Sum (insert a s) f = f a + Sum s f := Prod_insert_of_not_mem f H
theorem sum_union (f : A → B) {s₁ s₂ : finset A} (disj : s₁ ∩ s₂ = ∅) : theorem Sum_union (f : A → B) {s₁ s₂ : finset A} (disj : s₁ ∩ s₂ = ∅) :
sum (s₁ s₂) f = sum s₁ f + sum s₂ f := prod_union f disj Sum (s₁ s₂) f = Sum s₁ f + Sum s₂ f := Prod_union f disj
end decidable_eq end decidable_eq
theorem sum_add (s : finset A) (f g : A → B) : theorem Sum_add (s : finset A) (f g : A → B) :
sum s (λx, f x + g x) = sum s f + sum s g := prod_mul s f g Sum s (λx, f x + g x) = Sum s f + Sum s g := Prod_mul s f g
end add_comm_monoid end add_comm_monoid
end finset
end algebra end algebra

View file

@ -6,95 +6,98 @@ Author: Jeremy Avigad
Finite products and sums on the natural numbers. Finite products and sums on the natural numbers.
-/ -/
import data.nat.basic data.nat.order algebra.group_bigops import data.nat.basic data.nat.order algebra.group_bigops
open list open list finset
namespace nat namespace nat
open [classes] algebra open [classes] algebra
local attribute nat.comm_semiring [instance] local attribute nat.comm_semiring [instance]
variable {A : Type} variables {A : Type} [deceqA : decidable_eq A]
/- list.prod and list.sum -/ /- Prodl -/
definition list.prod (l : list A) (f : A → nat) : nat := algebra.list.prod l f definition Prodl (l : list A) (f : A → nat) : nat := algebra.Prodl l f
notation `∏` binders `←` l, r:(scoped f, list.prod l f) := r notation `∏` binders `←` l, r:(scoped f, Prodl l f) := r
definition list.sum (l : list A) (f : A → nat) : nat := algebra.list.sum l f
notation `∑` binders `←` l, r:(scoped f, list.sum l f) := r
namespace list -- i.e. nat.list theorem Prodl_nil (f : A → nat) : Prodl [] f = 1 := algebra.Prodl_nil f
open list -- i.e. ordinary lists theorem Prodl_cons (f : A → nat) (a : A) (l : list A) : Prodl (a::l) f = f a * Prodl l f :=
algebra.Prodl_cons f a l
theorem Prodl_append (l₁ l₂ : list A) (f : A → nat) : Prodl (l₁++l₂) f = Prodl l₁ f * Prodl l₂ f :=
algebra.Prodl_append l₁ l₂ f
theorem prod_nil (f : A → nat) : prod [] f = 1 := algebra.list.prod_nil f section deceqA
theorem prod_cons (f : A → nat) (a : A) (l : list A) : prod (a::l) f = f a * prod l f := include deceqA
algebra.list.prod_cons f a l theorem Prodl_insert_of_mem (f : A → nat) {a : A} {l : list A} (H : a ∈ l) :
theorem prod_append (l₁ l₂ : list A) (f : A → nat) : prod (l₁++l₂) f = prod l₁ f * prod l₂ f := Prodl (insert a l) f = Prodl l f := algebra.Prodl_insert_of_mem f H
algebra.list.prod_append l₁ l₂ f theorem Prodl_insert_of_not_mem (f : A → nat) {a : A} {l : list A} (H : a ∉ l) :
section decidable_eq Prodl (insert a l) f = f a * Prodl l f := algebra.Prodl_insert_of_not_mem f H
variable [H : decidable_eq A] theorem Prodl_union {l₁ l₂ : list A} (f : A → nat) (d : disjoint l₁ l₂) :
include H Prodl (union l₁ l₂) f = Prodl l₁ f * Prodl l₂ f := algebra.Prodl_union f d
theorem prod_insert_of_mem (f : A → nat) {a : A} {l : list A} (H : a ∈ l) : end deceqA
prod (insert a l) f = prod l f := algebra.list.prod_insert_of_mem f H
theorem prod_insert_of_not_mem (f : A → nat) {a : A} {l : list A} (H : a ∉ l) :
prod (insert a l) f = f a * prod l f := algebra.list.prod_insert_of_not_mem f H
theorem prod_union {l₁ l₂ : list A} (f : A → nat) (d : disjoint l₁ l₂) :
prod (union l₁ l₂) f = prod l₁ f * prod l₂ f := algebra.list.prod_union f d
end decidable_eq
theorem prod_mul (l : list A) (f g : A → nat) :
prod l (λx, f x * g x) = prod l f * prod l g := algebra.list.prod_mul l f g
theorem sum_nil (f : A → nat) : sum [] f = 0 := algebra.list.sum_nil f theorem Prodl_mul (l : list A) (f g : A → nat) :
theorem sum_cons (f : A → nat) (a : A) (l : list A) : sum (a::l) f = f a + sum l f := Prodl l (λx, f x * g x) = Prodl l f * Prodl l g := algebra.Prodl_mul l f g
algebra.list.sum_cons f a l
theorem sum_append (l₁ l₂ : list A) (f : A → nat) : sum (l₁++l₂) f = sum l₁ f + sum l₂ f :=
algebra.list.sum_append l₁ l₂ f
section decidable_eq
variable [H : decidable_eq A]
include H
theorem sum_insert_of_mem (f : A → nat) {a : A} {l : list A} (H : a ∈ l) :
sum (insert a l) f = sum l f := algebra.list.sum_insert_of_mem f H
theorem sum_insert_of_not_mem (f : A → nat) {a : A} {l : list A} (H : a ∉ l) :
sum (insert a l) f = f a + sum l f := algebra.list.sum_insert_of_not_mem f H
theorem sum_union {l₁ l₂ : list A} (f : A → nat) (d : disjoint l₁ l₂) :
sum (union l₁ l₂) f = sum l₁ f + sum l₂ f := algebra.list.sum_union f d
end decidable_eq
theorem sum_add (l : list A) (f g : A → nat) : sum l (λx, f x + g x) = sum l f + sum l g :=
algebra.list.sum_add l f g
end list
/- finset.prod and finset.sum -/ /- Prod -/
definition finset.prod (s : finset A) (f : A → nat) : nat := algebra.finset.prod s f definition Prod (s : finset A) (f : A → nat) : nat := algebra.Prod s f
notation `∏` binders `∈` s, r:(scoped f, finset.prod s f) := r notation `∏` binders `∈` s, r:(scoped f, Prod s f) := r
definition finset.sum (s : finset A) (f : A → nat) : nat := algebra.finset.sum s f
notation `∑` binders `∈` s, r:(scoped f, finset.sum s f) := r
namespace finset theorem Prod_empty (f : A → nat) : Prod ∅ f = 1 := algebra.Prod_empty f
open finset
theorem prod_empty (f : A → nat) : finset.prod ∅ f = 1 := algebra.finset.prod_empty f section deceqA
section decidable_eq include deceqA
variable [H : decidable_eq A] theorem Prod_insert_of_mem (f : A → nat) {a : A} {s : finset A} (H : a ∈ s) :
include H Prod (insert a s) f = Prod s f := algebra.Prod_insert_of_mem f H
theorem prod_insert_of_mem (f : A → nat) {a : A} {s : finset A} (H : a ∈ s) : theorem Prod_insert_of_not_mem (f : A → nat) {a : A} {s : finset A} (H : a ∉ s) :
prod (insert a s) f = prod s f := algebra.finset.prod_insert_of_mem f H Prod (insert a s) f = f a * Prod s f := algebra.Prod_insert_of_not_mem f H
theorem prod_insert_of_not_mem (f : A → nat) {a : A} {s : finset A} (H : a ∉ s) : theorem Prod_union (f : A → nat) {s₁ s₂ : finset A} (disj : s₁ ∩ s₂ = ∅) :
prod (insert a s) f = f a * prod s f := algebra.finset.prod_insert_of_not_mem f H Prod (s₁ s₂) f = Prod s₁ f * Prod s₂ f := algebra.Prod_union f disj
theorem prod_union (f : A → nat) {s₁ s₂ : finset A} (disj : s₁ ∩ s₂ = ∅) : end deceqA
prod (s₁ s₂) f = prod s₁ f * prod s₂ f := algebra.finset.prod_union f disj
end decidable_eq theorem Prod_mul (s : finset A) (f g : A → nat) : Prod s (λx, f x * g x) = Prod s f * Prod s g :=
theorem prod_mul (s : finset A) (f g : A → nat) : prod s (λx, f x * g x) = prod s f * prod s g := algebra.Prod_mul s f g
algebra.finset.prod_mul s f g
/- Suml -/
definition Suml (l : list A) (f : A → nat) : nat := algebra.Suml l f
notation `∑` binders `←` l, r:(scoped f, Suml l f) := r
theorem Suml_nil (f : A → nat) : Suml [] f = 0 := algebra.Suml_nil f
theorem Suml_cons (f : A → nat) (a : A) (l : list A) : Suml (a::l) f = f a + Suml l f :=
algebra.Suml_cons f a l
theorem Suml_append (l₁ l₂ : list A) (f : A → nat) : Suml (l₁++l₂) f = Suml l₁ f + Suml l₂ f :=
algebra.Suml_append l₁ l₂ f
section deceqA
include deceqA
theorem Suml_insert_of_mem (f : A → nat) {a : A} {l : list A} (H : a ∈ l) :
Suml (insert a l) f = Suml l f := algebra.Suml_insert_of_mem f H
theorem Suml_insert_of_not_mem (f : A → nat) {a : A} {l : list A} (H : a ∉ l) :
Suml (insert a l) f = f a + Suml l f := algebra.Suml_insert_of_not_mem f H
theorem Suml_union {l₁ l₂ : list A} (f : A → nat) (d : disjoint l₁ l₂) :
Suml (union l₁ l₂) f = Suml l₁ f + Suml l₂ f := algebra.Suml_union f d
end deceqA
theorem Suml_add (l : list A) (f g : A → nat) : Suml l (λx, f x + g x) = Suml l f + Suml l g :=
algebra.Suml_add l f g
/- Sum -/
definition Sum (s : finset A) (f : A → nat) : nat := algebra.Sum s f
notation `∑` binders `∈` s, r:(scoped f, Sum s f) := r
theorem Sum_empty (f : A → nat) : Sum ∅ f = 0 := algebra.Sum_empty f
section deceqA
include deceqA
theorem Sum_insert_of_mem (f : A → nat) {a : A} {s : finset A} (H : a ∈ s) :
Sum (insert a s) f = Sum s f := algebra.Sum_insert_of_mem f H
theorem Sum_insert_of_not_mem (f : A → nat) {a : A} {s : finset A} (H : a ∉ s) :
Sum (insert a s) f = f a + Sum s f := algebra.Sum_insert_of_not_mem f H
theorem Sum_union (f : A → nat) {s₁ s₂ : finset A} (disj : s₁ ∩ s₂ = ∅) :
Sum (s₁ s₂) f = Sum s₁ f + Sum s₂ f := algebra.Sum_union f disj
end deceqA
theorem Sum_add (s : finset A) (f g : A → nat) : Sum s (λx, f x + g x) = Sum s f + Sum s g :=
algebra.Sum_add s f g
theorem sum_empty (f : A → nat) : finset.sum ∅ f = 0 := algebra.finset.sum_empty f
section decidable_eq
variable [H : decidable_eq A]
include H
theorem sum_insert_of_mem (f : A → nat) {a : A} {s : finset A} (H : a ∈ s) :
sum (insert a s) f = sum s f := algebra.finset.sum_insert_of_mem f H
theorem sum_insert_of_not_mem (f : A → nat) {a : A} {s : finset A} (H : a ∉ s) :
sum (insert a s) f = f a + sum s f := algebra.finset.sum_insert_of_not_mem f H
theorem sum_union (f : A → nat) {s₁ s₂ : finset A} (disj : s₁ ∩ s₂ = ∅) :
sum (s₁ s₂) f = sum s₁ f + sum s₂ f := algebra.finset.sum_union f disj
end decidable_eq
theorem sum_add (s : finset A) (f g : A → nat) : sum s (λx, f x + g x) = sum s f + sum s g :=
algebra.finset.sum_add s f g
end finset
end nat end nat