From 5816424b5cf35e552c7d4d901c46aaa0b8461c03 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Sun, 23 Jul 2017 18:26:42 +0100 Subject: [PATCH] refactor(*): use 'lemma' iff statement is private --- algebra/lattice/basic.lean | 108 ++--- algebra/lattice/boolean_algebra.lean | 50 +-- algebra/lattice/bounded_lattice.lean | 4 +- algebra/lattice/complete_boolean_algebra.lean | 12 +- algebra/lattice/complete_lattice.lean | 194 ++++----- algebra/lattice/filter.lean | 396 +++++++++--------- algebra/lattice/fixed_points.lean | 30 +- algebra/lattice/zorn.lean | 28 +- algebra/order.lean | 14 +- data/bitvec.lean | 4 +- data/bool.lean | 24 +- data/fin.lean | 4 +- data/hash_map.lean | 14 +- data/int/basic.lean | 2 +- data/list/basic.lean | 66 +-- data/list/comb.lean | 16 +- data/list/perm.lean | 8 +- data/list/set.lean | 54 +-- data/list/sort.lean | 4 +- data/nat/basic.lean | 8 +- data/nat/gcd.lean | 2 +- data/nat/sub.lean | 6 +- data/num/lemmas.lean | 22 +- data/rat.lean | 40 +- data/seq/computation.lean | 22 +- data/seq/parallel.lean | 6 +- data/seq/seq.lean | 50 +-- data/seq/wseq.lean | 28 +- data/set/basic.lean | 2 +- data/set/finite.lean | 12 +- data/stream.lean | 220 +++++----- data/vector.lean | 20 +- logic/basic.lean | 20 +- tactic/converter/binders.lean | 18 +- tests/examples.lean | 2 +- tests/finish2.lean | 2 +- topology/continuity.lean | 92 ++-- topology/topological_space.lean | 130 +++--- topology/uniform_space.lean | 128 +++--- 39 files changed, 931 insertions(+), 931 deletions(-) diff --git a/algebra/lattice/basic.lean b/algebra/lattice/basic.lean index 2cbd1968a9d96..0da66876819c4 100644 --- a/algebra/lattice/basic.lean +++ b/algebra/lattice/basic.lean @@ -20,7 +20,7 @@ section variable {α : Type u} -- TODO: this seems crazy, but it also seems to work reasonably well - @[ematch] lemma le_antisymm' [weak_order α] : ∀ {a b : α}, (: a ≤ b :) → b ≤ a → a = b := + @[ematch] theorem le_antisymm' [weak_order α] : ∀ {a b : α}, (: a ≤ b :) → b ≤ a → a = b := weak_order.le_antisymm end @@ -47,17 +47,17 @@ class order_top (α : Type u) extends has_top α, weak_order α := section order_top variables {α : Type u} [order_top α] {a : α} -@[simp] lemma le_top : a ≤ ⊤ := +@[simp] theorem le_top : a ≤ ⊤ := order_top.le_top a -lemma top_unique (h : ⊤ ≤ a) : a = ⊤ := +theorem top_unique (h : ⊤ ≤ a) : a = ⊤ := le_antisymm le_top h -- TODO: delete in favor of the next? -lemma eq_top_iff : a = ⊤ ↔ ⊤ ≤ a := +theorem eq_top_iff : a = ⊤ ↔ ⊤ ≤ a := ⟨assume eq, eq.symm ▸ le_refl ⊤, top_unique⟩ -@[simp] lemma top_le_iff : ⊤ ≤ a ↔ a = ⊤ := +@[simp] theorem top_le_iff : ⊤ ≤ a ↔ a = ⊤ := ⟨top_unique, λ h, h.symm ▸ le_refl ⊤⟩ end order_top @@ -68,19 +68,19 @@ class order_bot (α : Type u) extends has_bot α, weak_order α := section order_bot variables {α : Type u} [order_bot α] {a : α} -@[simp] lemma bot_le : ⊥ ≤ a := order_bot.bot_le a +@[simp] theorem bot_le : ⊥ ≤ a := order_bot.bot_le a -lemma bot_unique (h : a ≤ ⊥) : a = ⊥ := +theorem bot_unique (h : a ≤ ⊥) : a = ⊥ := le_antisymm h bot_le -- TODO: delete? -lemma eq_bot_iff : a = ⊥ ↔ a ≤ ⊥ := +theorem eq_bot_iff : a = ⊥ ↔ a ≤ ⊥ := ⟨assume eq, eq.symm ▸ le_refl ⊥, bot_unique⟩ -@[simp] lemma le_bot_iff : a ≤ ⊥ ↔ a = ⊥ := +@[simp] theorem le_bot_iff : a ≤ ⊥ ↔ a = ⊥ := ⟨bot_unique, assume h, h.symm ▸ le_refl ⊥⟩ -lemma neq_bot_of_le_neq_bot {a b : α} (hb : b ≠ ⊥) (hab : b ≤ a) : a ≠ ⊥ := +theorem neq_bot_of_le_neq_bot {a b : α} (hb : b ≠ ⊥) (hab : b ≤ a) : a ≠ ⊥ := assume ha, hb $ bot_unique $ ha ▸ hab end order_bot @@ -93,55 +93,55 @@ class semilattice_sup (α : Type u) extends has_sup α, weak_order α := section semilattice_sup variables {α : Type u} [semilattice_sup α] {a b c d : α} -lemma le_sup_left : a ≤ a ⊔ b := +theorem le_sup_left : a ≤ a ⊔ b := semilattice_sup.le_sup_left a b -@[ematch] lemma le_sup_left' : a ≤ (: a ⊔ b :) := +@[ematch] theorem le_sup_left' : a ≤ (: a ⊔ b :) := semilattice_sup.le_sup_left a b -lemma le_sup_right : b ≤ a ⊔ b := +theorem le_sup_right : b ≤ a ⊔ b := semilattice_sup.le_sup_right a b -@[ematch] lemma le_sup_right' : b ≤ (: a ⊔ b :) := +@[ematch] theorem le_sup_right' : b ≤ (: a ⊔ b :) := semilattice_sup.le_sup_right a b -lemma le_sup_left_of_le (h : c ≤ a) : c ≤ a ⊔ b := +theorem le_sup_left_of_le (h : c ≤ a) : c ≤ a ⊔ b := by finish -lemma le_sup_right_of_le (h : c ≤ b) : c ≤ a ⊔ b := +theorem le_sup_right_of_le (h : c ≤ b) : c ≤ a ⊔ b := by finish -lemma sup_le : a ≤ c → b ≤ c → a ⊔ b ≤ c := +theorem sup_le : a ≤ c → b ≤ c → a ⊔ b ≤ c := semilattice_sup.sup_le a b c -@[simp] lemma sup_le_iff : a ⊔ b ≤ c ↔ a ≤ c ∧ b ≤ c := +@[simp] theorem sup_le_iff : a ⊔ b ≤ c ↔ a ≤ c ∧ b ≤ c := ⟨assume h : a ⊔ b ≤ c, ⟨le_trans le_sup_left h, le_trans le_sup_right h⟩, assume ⟨h₁, h₂⟩, sup_le h₁ h₂⟩ -- TODO: if we just write le_antisymm, Lean doesn't know which ≤ we want to use -- Can we do anything about that? -lemma sup_of_le_left (h : b ≤ a) : a ⊔ b = a := +theorem sup_of_le_left (h : b ≤ a) : a ⊔ b = a := by apply le_antisymm; finish -lemma sup_of_le_right (h : a ≤ b) : a ⊔ b = b := +theorem sup_of_le_right (h : a ≤ b) : a ⊔ b = b := by apply le_antisymm; finish -lemma sup_le_sup (h₁ : a ≤ b) (h₂ : c ≤ d) : a ⊔ c ≤ b ⊔ d := +theorem sup_le_sup (h₁ : a ≤ b) (h₂ : c ≤ d) : a ⊔ c ≤ b ⊔ d := by finish -lemma le_of_sup_eq (h : a ⊔ b = b) : a ≤ b := +theorem le_of_sup_eq (h : a ⊔ b = b) : a ≤ b := by finish -@[simp] lemma sup_idem : a ⊔ a = a := +@[simp] theorem sup_idem : a ⊔ a = a := by apply le_antisymm; finish -lemma sup_comm : a ⊔ b = b ⊔ a := +theorem sup_comm : a ⊔ b = b ⊔ a := by apply le_antisymm; finish instance semilattice_sup_to_is_commutative [semilattice_sup α] : is_commutative α (⊔) := ⟨@sup_comm _ _⟩ -lemma sup_assoc : a ⊔ b ⊔ c = a ⊔ (b ⊔ c) := +theorem sup_assoc : a ⊔ b ⊔ c = a ⊔ (b ⊔ c) := by apply le_antisymm; finish instance semilattice_sup_to_is_associative [semilattice_sup α] : is_associative α (⊔) := @@ -157,53 +157,53 @@ class semilattice_inf (α : Type u) extends has_inf α, weak_order α := section semilattice_inf variables {α : Type u} [semilattice_inf α] {a b c d : α} -lemma inf_le_left : a ⊓ b ≤ a := +theorem inf_le_left : a ⊓ b ≤ a := semilattice_inf.inf_le_left a b -@[ematch] lemma inf_le_left' : (: a ⊓ b :) ≤ a := +@[ematch] theorem inf_le_left' : (: a ⊓ b :) ≤ a := semilattice_inf.inf_le_left a b -lemma inf_le_right : a ⊓ b ≤ b := +theorem inf_le_right : a ⊓ b ≤ b := semilattice_inf.inf_le_right a b -@[ematch] lemma inf_le_right' : (: a ⊓ b :) ≤ b := +@[ematch] theorem inf_le_right' : (: a ⊓ b :) ≤ b := semilattice_inf.inf_le_right a b -lemma le_inf : a ≤ b → a ≤ c → a ≤ b ⊓ c := +theorem le_inf : a ≤ b → a ≤ c → a ≤ b ⊓ c := semilattice_inf.le_inf a b c -lemma inf_le_left_of_le (h : a ≤ c) : a ⊓ b ≤ c := +theorem inf_le_left_of_le (h : a ≤ c) : a ⊓ b ≤ c := le_trans inf_le_left h -lemma inf_le_right_of_le (h : b ≤ c) : a ⊓ b ≤ c := +theorem inf_le_right_of_le (h : b ≤ c) : a ⊓ b ≤ c := le_trans inf_le_right h -@[simp] lemma le_inf_iff : a ≤ b ⊓ c ↔ a ≤ b ∧ a ≤ c := +@[simp] theorem le_inf_iff : a ≤ b ⊓ c ↔ a ≤ b ∧ a ≤ c := ⟨assume h : a ≤ b ⊓ c, ⟨le_trans h inf_le_left, le_trans h inf_le_right⟩, assume ⟨h₁, h₂⟩, le_inf h₁ h₂⟩ -lemma inf_of_le_left (h : a ≤ b) : a ⊓ b = a := +theorem inf_of_le_left (h : a ≤ b) : a ⊓ b = a := by apply le_antisymm; finish -lemma inf_of_le_right (h : b ≤ a) : a ⊓ b = b := +theorem inf_of_le_right (h : b ≤ a) : a ⊓ b = b := by apply le_antisymm; finish -lemma inf_le_inf (h₁ : a ≤ b) (h₂ : c ≤ d) : a ⊓ c ≤ b ⊓ d := +theorem inf_le_inf (h₁ : a ≤ b) (h₂ : c ≤ d) : a ⊓ c ≤ b ⊓ d := by finish -lemma le_of_inf_eq (h : a ⊓ b = a) : a ≤ b := +theorem le_of_inf_eq (h : a ⊓ b = a) : a ≤ b := by finish -@[simp] lemma inf_idem : a ⊓ a = a := +@[simp] theorem inf_idem : a ⊓ a = a := by apply le_antisymm; finish -lemma inf_comm : a ⊓ b = b ⊓ a := +theorem inf_comm : a ⊓ b = b ⊓ a := by apply le_antisymm; finish instance semilattice_inf_to_is_commutative [semilattice_inf α] : is_commutative α (⊓) := ⟨@inf_comm _ _⟩ -lemma inf_assoc : a ⊓ b ⊓ c = a ⊓ (b ⊓ c) := +theorem inf_assoc : a ⊓ b ⊓ c = a ⊓ (b ⊓ c) := by apply le_antisymm; finish instance semilattice_inf_to_is_associative [semilattice_inf α] : is_associative α (⊓) := @@ -216,10 +216,10 @@ class semilattice_sup_top (α : Type u) extends order_top α, semilattice_sup α section semilattice_sup_top variables {α : Type u} [semilattice_sup_top α] {a : α} -@[simp] lemma top_sup_eq : ⊤ ⊔ a = ⊤ := +@[simp] theorem top_sup_eq : ⊤ ⊔ a = ⊤ := sup_of_le_left le_top -@[simp] lemma sup_top_eq : a ⊔ ⊤ = ⊤ := +@[simp] theorem sup_top_eq : a ⊔ ⊤ = ⊤ := sup_of_le_right le_top end semilattice_sup_top @@ -229,13 +229,13 @@ class semilattice_sup_bot (α : Type u) extends order_bot α, semilattice_sup α section semilattice_sup_bot variables {α : Type u} [semilattice_sup_bot α] {a b : α} -@[simp] lemma bot_sup_eq : ⊥ ⊔ a = a := +@[simp] theorem bot_sup_eq : ⊥ ⊔ a = a := sup_of_le_right bot_le -@[simp] lemma sup_bot_eq : a ⊔ ⊥ = a := +@[simp] theorem sup_bot_eq : a ⊔ ⊥ = a := sup_of_le_left bot_le -@[simp] lemma sup_eq_bot_iff : a ⊔ b = ⊥ ↔ (a = ⊥ ∧ b = ⊥) := +@[simp] theorem sup_eq_bot_iff : a ⊔ b = ⊥ ↔ (a = ⊥ ∧ b = ⊥) := by rw [eq_bot_iff, sup_le_iff]; simp end semilattice_sup_bot @@ -245,13 +245,13 @@ class semilattice_inf_top (α : Type u) extends order_top α, semilattice_inf α section semilattice_inf_top variables {α : Type u} [semilattice_inf_top α] {a b : α} -@[simp] lemma top_inf_eq : ⊤ ⊓ a = a := +@[simp] theorem top_inf_eq : ⊤ ⊓ a = a := inf_of_le_right le_top -@[simp] lemma inf_top_eq : a ⊓ ⊤ = a := +@[simp] theorem inf_top_eq : a ⊓ ⊤ = a := inf_of_le_left le_top -@[simp] lemma inf_eq_top_iff : a ⊓ b = ⊤ ↔ (a = ⊤ ∧ b = ⊤) := +@[simp] theorem inf_eq_top_iff : a ⊓ b = ⊤ ↔ (a = ⊤ ∧ b = ⊤) := by rw [eq_top_iff, le_inf_iff]; simp end semilattice_inf_top @@ -261,10 +261,10 @@ class semilattice_inf_bot (α : Type u) extends order_bot α, semilattice_inf α section semilattice_inf_bot variables {α : Type u} [semilattice_inf_bot α] {a : α} -@[simp] lemma bot_inf_eq : ⊥ ⊓ a = ⊥ := +@[simp] theorem bot_inf_eq : ⊥ ⊓ a = ⊥ := inf_of_le_left bot_le -@[simp] lemma inf_bot_eq : a ⊓ ⊥ = ⊥ := +@[simp] theorem inf_bot_eq : a ⊓ ⊥ = ⊥ := inf_of_le_right bot_le end semilattice_inf_bot @@ -278,16 +278,16 @@ variables {α : Type u} [lattice α] {a b c d : α} /- Distributivity laws -/ /- TODO: better names? -/ -lemma sup_inf_le : a ⊔ (b ⊓ c) ≤ (a ⊔ b) ⊓ (a ⊔ c) := +theorem sup_inf_le : a ⊔ (b ⊓ c) ≤ (a ⊔ b) ⊓ (a ⊔ c) := by finish -lemma le_inf_sup : (a ⊓ b) ⊔ (a ⊓ c) ≤ a ⊓ (b ⊔ c) := +theorem le_inf_sup : (a ⊓ b) ⊔ (a ⊓ c) ≤ a ⊓ (b ⊔ c) := by finish -lemma inf_sup_self : a ⊓ (a ⊔ b) = a := +theorem inf_sup_self : a ⊓ (a ⊔ b) = a := le_antisymm (by finish) (by finish) -lemma sup_inf_self : a ⊔ (a ⊓ b) = a := +theorem sup_inf_self : a ⊔ (a ⊓ b) = a := le_antisymm (by finish) (by finish) end lattice diff --git a/algebra/lattice/boolean_algebra.lean b/algebra/lattice/boolean_algebra.lean index 29dd7f6183699..81603cc65f705 100644 --- a/algebra/lattice/boolean_algebra.lean +++ b/algebra/lattice/boolean_algebra.lean @@ -20,23 +20,23 @@ class distrib_lattice α extends lattice α := section distrib_lattice variables [distrib_lattice α] -lemma le_sup_inf : ∀{x y z : α}, (x ⊔ y) ⊓ (x ⊔ z) ≤ x ⊔ (y ⊓ z) := +theorem le_sup_inf : ∀{x y z : α}, (x ⊔ y) ⊓ (x ⊔ z) ≤ x ⊔ (y ⊓ z) := distrib_lattice.le_sup_inf -lemma sup_inf_left {x y z : α} : x ⊔ (y ⊓ z) = (x ⊔ y) ⊓ (x ⊔ z) := +theorem sup_inf_left {x y z : α} : x ⊔ (y ⊓ z) = (x ⊔ y) ⊓ (x ⊔ z) := le_antisymm sup_inf_le le_sup_inf -lemma sup_inf_right : (y ⊓ z) ⊔ x = (y ⊔ x) ⊓ (z ⊔ x) := +theorem sup_inf_right : (y ⊓ z) ⊔ x = (y ⊔ x) ⊓ (z ⊔ x) := by simp [sup_inf_left, λy:α, @sup_comm α _ y x] -lemma inf_sup_left : x ⊓ (y ⊔ z) = (x ⊓ y) ⊔ (x ⊓ z) := +theorem inf_sup_left : x ⊓ (y ⊔ z) = (x ⊓ y) ⊔ (x ⊓ z) := calc x ⊓ (y ⊔ z) = (x ⊓ (x ⊔ z)) ⊓ (y ⊔ z) : by rw [inf_sup_self] ... = x ⊓ ((x ⊓ y) ⊔ z) : by simp [inf_assoc, sup_inf_right] ... = (x ⊔ (x ⊓ y)) ⊓ ((x ⊓ y) ⊔ z) : by rw [sup_inf_self] ... = ((x ⊓ y) ⊔ x) ⊓ ((x ⊓ y) ⊔ z) : by rw [sup_comm] ... = (x ⊓ y) ⊔ (x ⊓ z) : by rw [sup_inf_left] -lemma inf_sup_right : (y ⊔ z) ⊓ x = (y ⊓ x) ⊔ (z ⊓ x) := +theorem inf_sup_right : (y ⊔ z) ⊓ x = (y ⊓ x) ⊔ (z ⊓ x) := by simp [inf_sup_left, λy:α, @inf_comm α _ y x] end distrib_lattice @@ -52,25 +52,25 @@ section boolean_algebra variables [boolean_algebra α] @[simp] -lemma inf_neg_eq_bot : x ⊓ - x = ⊥ := +theorem inf_neg_eq_bot : x ⊓ - x = ⊥ := boolean_algebra.inf_neg_eq_bot x @[simp] -lemma neg_inf_eq_bot : - x ⊓ x = ⊥ := +theorem neg_inf_eq_bot : - x ⊓ x = ⊥ := eq.trans inf_comm inf_neg_eq_bot @[simp] -lemma sup_neg_eq_top : x ⊔ - x = ⊤ := +theorem sup_neg_eq_top : x ⊔ - x = ⊤ := boolean_algebra.sup_neg_eq_top x @[simp] -lemma neg_sup_eq_top : - x ⊔ x = ⊤ := +theorem neg_sup_eq_top : - x ⊔ x = ⊤ := eq.trans sup_comm sup_neg_eq_top -lemma sub_eq : x - y = x ⊓ - y := +theorem sub_eq : x - y = x ⊓ - y := boolean_algebra.sub_eq x y -lemma neg_unique (i : x ⊓ y = ⊥) (s : x ⊔ y = ⊤) : - x = y := +theorem neg_unique (i : x ⊓ y = ⊥) (s : x ⊔ y = ⊤) : - x = y := have (- x ⊓ x) ⊔ (- x ⊓ y) = (y ⊓ x) ⊔ (y ⊓ - x), by rsimp, have - x ⊓ (x ⊔ y) = y ⊓ (x ⊔ - x), @@ -78,28 +78,28 @@ have - x ⊓ (x ⊔ y) = y ⊓ (x ⊔ - x), by rsimp @[simp] -lemma neg_top : - ⊤ = (⊥:α) := +theorem neg_top : - ⊤ = (⊥:α) := neg_unique (by simp) (by simp) @[simp] -lemma neg_bot : - ⊥ = (⊤:α) := +theorem neg_bot : - ⊥ = (⊤:α) := neg_unique (by simp) (by simp) @[simp] -lemma neg_neg : - (- x) = x := +theorem neg_neg : - (- x) = x := neg_unique (by simp) (by simp) -lemma neg_eq_neg_of_eq (h : - x = - y) : x = y := +theorem neg_eq_neg_of_eq (h : - x = - y) : x = y := have - - x = - - y, from congr_arg has_neg.neg h, by simp [neg_neg] at this; assumption @[simp] -lemma neg_eq_neg_iff : - x = - y ↔ x = y := +theorem neg_eq_neg_iff : - x = - y ↔ x = y := ⟨neg_eq_neg_of_eq, congr_arg has_neg.neg⟩ @[simp] -lemma neg_inf : - (x ⊓ y) = -x ⊔ -y := +theorem neg_inf : - (x ⊓ y) = -x ⊔ -y := neg_unique -- TODO: try rsimp if it supports custom lemmas (calc (x ⊓ y) ⊓ (- x ⊔ - y) = (y ⊓ (x ⊓ - x)) ⊔ (x ⊓ (y ⊓ - y)) : by rw [inf_sup_left]; ac_refl ... = ⊥ : by simp) @@ -107,33 +107,33 @@ neg_unique -- TODO: try rsimp if it supports custom lemmas ... = ⊤ : by simp) @[simp] -lemma neg_sup : - (x ⊔ y) = -x ⊓ -y := +theorem neg_sup : - (x ⊔ y) = -x ⊓ -y := begin [smt] eblast_using [neg_neg, neg_inf] end -lemma neg_le_neg (h : y ≤ x) : - x ≤ - y := +theorem neg_le_neg (h : y ≤ x) : - x ≤ - y := le_of_inf_eq $ calc -x ⊓ -y = - (x ⊔ y) : neg_sup.symm ... = -x : congr_arg has_neg.neg $ sup_of_le_left h -lemma neg_le_neg_iff_le : - y ≤ - x ↔ x ≤ y := +theorem neg_le_neg_iff_le : - y ≤ - x ↔ x ≤ y := ⟨assume h, by have h := neg_le_neg h; simp at h; assumption, neg_le_neg⟩ -lemma le_neg_of_le_neg (h : y ≤ - x) : x ≤ - y := +theorem le_neg_of_le_neg (h : y ≤ - x) : x ≤ - y := have - (- x) ≤ - y, from neg_le_neg h, by simp at this; assumption -lemma neg_le_of_neg_le (h : - y ≤ x) : - x ≤ y := +theorem neg_le_of_neg_le (h : - y ≤ x) : - x ≤ y := have - x ≤ - (- y), from neg_le_neg h, by simp at this; assumption -lemma neg_le_iff_neg_le : y ≤ - x ↔ x ≤ - y := +theorem neg_le_iff_neg_le : y ≤ - x ↔ x ≤ - y := ⟨le_neg_of_le_neg, le_neg_of_le_neg⟩ -lemma sup_sub_same : x ⊔ (y - x) = x ⊔ y := +theorem sup_sub_same : x ⊔ (y - x) = x ⊔ y := by simp [sub_eq, sup_inf_left] -lemma sub_eq_left (h : x ⊓ y = ⊥) : x - y = x := +theorem sub_eq_left (h : x ⊓ y = ⊥) : x - y = x := calc x - y = (x ⊓ -y) ⊔ (x ⊓ y) : by simp [h, sub_eq] ... = (-y ⊓ x) ⊔ (y ⊓ x) : by simp [inf_comm] ... = (-y ⊔ y) ⊓ x : inf_sup_right.symm diff --git a/algebra/lattice/bounded_lattice.lean b/algebra/lattice/bounded_lattice.lean index 348c710b84cde..5386e8324c1ed 100644 --- a/algebra/lattice/bounded_lattice.lean +++ b/algebra/lattice/bounded_lattice.lean @@ -60,12 +60,12 @@ instance bounded_lattice_Prop : bounded_lattice Prop := section logic variable [weak_order α] -lemma monotone_and {p q : α → Prop} (m_p : monotone p) (m_q : monotone q) : +theorem monotone_and {p q : α → Prop} (m_p : monotone p) (m_q : monotone q) : monotone (λx, p x ∧ q x) := assume a b h, and.imp (m_p h) (m_q h) -- Note: by finish [monotone] doesn't work -lemma monotone_or {p q : α → Prop} (m_p : monotone p) (m_q : monotone q) : +theorem monotone_or {p q : α → Prop} (m_p : monotone p) (m_q : monotone q) : monotone (λx, p x ∨ q x) := assume a b h, or.imp (m_p h) (m_q h) end logic diff --git a/algebra/lattice/complete_boolean_algebra.lean b/algebra/lattice/complete_boolean_algebra.lean index 5d2b17f5579ea..0bc329c349584 100644 --- a/algebra/lattice/complete_boolean_algebra.lean +++ b/algebra/lattice/complete_boolean_algebra.lean @@ -21,12 +21,12 @@ class complete_distrib_lattice α extends complete_lattice α := section complete_distrib_lattice variables [complete_distrib_lattice α] {a : α} {s : set α} -lemma sup_Inf_eq : a ⊔ Inf s = (⨅ b ∈ s, a ⊔ b) := +theorem sup_Inf_eq : a ⊔ Inf s = (⨅ b ∈ s, a ⊔ b) := le_antisymm (le_infi $ assume i, le_infi $ assume h, sup_le_sup (le_refl _) (Inf_le h)) (complete_distrib_lattice.infi_sup_le_sup_Inf _ _) -lemma inf_Sup_eq : a ⊓ Sup s = (⨆ b ∈ s, a ⊓ b) := +theorem inf_Sup_eq : a ⊓ Sup s = (⨆ b ∈ s, a ⊓ b) := le_antisymm (complete_distrib_lattice.inf_Sup_le_supr_inf _ _) (supr_le $ assume i, supr_le $ assume h, inf_le_inf (le_refl _) (le_Sup h)) @@ -45,18 +45,18 @@ class complete_boolean_algebra α extends boolean_algebra α, complete_distrib_l section complete_boolean_algebra variables [complete_boolean_algebra α] {a b : α} {s : set α} {f : ι → α} -lemma neg_infi : - infi f = (⨆i, - f i) := +theorem neg_infi : - infi f = (⨆i, - f i) := le_antisymm (neg_le_of_neg_le $ le_infi $ assume i, neg_le_of_neg_le $ le_supr (λi, - f i) i) (supr_le $ assume i, neg_le_neg $ infi_le _ _) -lemma neg_supr : - supr f = (⨅i, - f i) := +theorem neg_supr : - supr f = (⨅i, - f i) := neg_eq_neg_of_eq (by simp [neg_infi]) -lemma neg_Inf : - Inf s = (⨆i∈s, - i) := +theorem neg_Inf : - Inf s = (⨆i∈s, - i) := by simp [Inf_eq_infi, neg_infi] -lemma neg_Sup : - Sup s = (⨅i∈s, - i) := +theorem neg_Sup : - Sup s = (⨅i∈s, - i) := by simp [Sup_eq_supr, neg_supr] end complete_boolean_algebra diff --git a/algebra/lattice/complete_lattice.lean b/algebra/lattice/complete_lattice.lean index ac85676f97b83..c92bde5894d8e 100644 --- a/algebra/lattice/complete_lattice.lean +++ b/algebra/lattice/complete_lattice.lean @@ -36,46 +36,46 @@ open set variables [complete_lattice α] {s t : set α} {a b : α} @[ematch] -lemma le_Sup : a ∈ s → a ≤ Sup s := complete_lattice.le_Sup s a +theorem le_Sup : a ∈ s → a ≤ Sup s := complete_lattice.le_Sup s a -lemma Sup_le : (∀b∈s, b ≤ a) → Sup s ≤ a := complete_lattice.Sup_le s a +theorem Sup_le : (∀b∈s, b ≤ a) → Sup s ≤ a := complete_lattice.Sup_le s a @[ematch] -lemma Inf_le : a ∈ s → Inf s ≤ a := complete_lattice.Inf_le s a +theorem Inf_le : a ∈ s → Inf s ≤ a := complete_lattice.Inf_le s a -lemma le_Inf : (∀b∈s, a ≤ b) → a ≤ Inf s := complete_lattice.le_Inf s a +theorem le_Inf : (∀b∈s, a ≤ b) → a ≤ Inf s := complete_lattice.le_Inf s a -lemma le_Sup_of_le (hb : b ∈ s) (h : a ≤ b) : a ≤ Sup s := +theorem le_Sup_of_le (hb : b ∈ s) (h : a ≤ b) : a ≤ Sup s := le_trans h (le_Sup hb) -lemma Inf_le_of_le (hb : b ∈ s) (h : b ≤ a) : Inf s ≤ a := +theorem Inf_le_of_le (hb : b ∈ s) (h : b ≤ a) : Inf s ≤ a := le_trans (Inf_le hb) h -lemma Sup_le_Sup (h : s ⊆ t) : Sup s ≤ Sup t := +theorem Sup_le_Sup (h : s ⊆ t) : Sup s ≤ Sup t := Sup_le (assume a, assume ha : a ∈ s, le_Sup $ h ha) -lemma Inf_le_Inf (h : s ⊆ t) : Inf t ≤ Inf s := +theorem Inf_le_Inf (h : s ⊆ t) : Inf t ≤ Inf s := le_Inf (assume a, assume ha : a ∈ s, Inf_le $ h ha) @[simp] -lemma le_Sup_iff : Sup s ≤ a ↔ (∀b ∈ s, b ≤ a) := +theorem le_Sup_iff : Sup s ≤ a ↔ (∀b ∈ s, b ≤ a) := ⟨assume : Sup s ≤ a, assume b, assume : b ∈ s, le_trans (le_Sup ‹b ∈ s›) ‹Sup s ≤ a›, Sup_le⟩ @[simp] -lemma Inf_le_iff : a ≤ Inf s ↔ (∀b ∈ s, a ≤ b) := +theorem Inf_le_iff : a ≤ Inf s ↔ (∀b ∈ s, a ≤ b) := ⟨assume : a ≤ Inf s, assume b, assume : b ∈ s, le_trans ‹a ≤ Inf s› (Inf_le ‹b ∈ s›), le_Inf⟩ -- how to state this? instead a parameter `a`, use `∃a, a ∈ s` or `s ≠ ∅`? -lemma Inf_le_Sup (h : a ∈ s) : Inf s ≤ Sup s := +theorem Inf_le_Sup (h : a ∈ s) : Inf s ≤ Sup s := by have := le_Sup h; finish --Inf_le_of_le h (le_Sup h) -- TODO: it is weird that we have to add union_def -lemma Sup_union {s t : set α} : Sup (s ∪ t) = Sup s ⊔ Sup t := +theorem Sup_union {s t : set α} : Sup (s ∪ t) = Sup s ⊔ Sup t := le_antisymm (by finish) (by finish [union_def]) /- old proof: @@ -84,13 +84,13 @@ le_antisymm (sup_le (Sup_le_Sup $ subset_union_left _ _) (Sup_le_Sup $ subset_union_right _ _)) -/ -lemma Sup_inter_le {s t : set α} : Sup (s ∩ t) ≤ Sup s ⊓ Sup t := +theorem Sup_inter_le {s t : set α} : Sup (s ∩ t) ≤ Sup s ⊓ Sup t := by finish /- Sup_le (assume a ⟨a_s, a_t⟩, le_inf (le_Sup a_s) (le_Sup a_t)) -/ -lemma Inf_union {s t : set α} : Inf (s ∪ t) = Inf s ⊓ Inf t := +theorem Inf_union {s t : set α} : Inf (s ∪ t) = Inf s ⊓ Inf t := le_antisymm (by finish [union_def]) (by finish) /- old proof: @@ -99,33 +99,33 @@ le_antisymm (le_Inf $ assume a h, or.rec_on h (inf_le_left_of_le ∘ Inf_le) (inf_le_right_of_le ∘ Inf_le)) -/ -lemma le_Inf_inter {s t : set α} : Inf s ⊔ Inf t ≤ Inf (s ∩ t) := +theorem le_Inf_inter {s t : set α} : Inf s ⊔ Inf t ≤ Inf (s ∩ t) := by finish /- le_Inf (assume a ⟨a_s, a_t⟩, sup_le (Inf_le a_s) (Inf_le a_t)) -/ @[simp] -lemma Sup_empty : Sup ∅ = (⊥ : α) := +theorem Sup_empty : Sup ∅ = (⊥ : α) := le_antisymm (by finish) (by finish) -- le_antisymm (Sup_le (assume _, false.elim)) bot_le @[simp] -lemma Inf_empty : Inf ∅ = (⊤ : α) := +theorem Inf_empty : Inf ∅ = (⊤ : α) := le_antisymm (by finish) (by finish) --le_antisymm le_top (le_Inf (assume _, false.elim)) @[simp] -lemma Sup_univ : Sup univ = (⊤ : α) := +theorem Sup_univ : Sup univ = (⊤ : α) := le_antisymm (by finish) (le_Sup ⟨⟩) -- finish fails because ⊤ ≤ a simplifies to a = ⊤ --le_antisymm le_top (le_Sup ⟨⟩) @[simp] -lemma Inf_univ : Inf univ = (⊥ : α) := +theorem Inf_univ : Inf univ = (⊥ : α) := le_antisymm (Inf_le ⟨⟩) bot_le @[simp] -lemma Sup_insert {a : α} {s : set α} : Sup (insert a s) = a ⊔ Sup s := +theorem Sup_insert {a : α} {s : set α} : Sup (insert a s) = a ⊔ Sup s := le_antisymm (by finish) (by finish [insert_def]) /- old proof @@ -136,7 +136,7 @@ calc Sup (insert a s) = Sup {b | b = a} ⊔ Sup s : Sup_union -/ @[simp] -lemma Inf_insert {a : α} {s : set α} : Inf (insert a s) = a ⊓ Inf s := +theorem Inf_insert {a : α} {s : set α} : Inf (insert a s) = a ⊓ Inf s := le_antisymm (by finish [insert_def]) (by finish) /- old proof @@ -147,12 +147,12 @@ calc Inf (insert a s) = Inf {b | b = a} ⊓ Inf s : Inf_union -/ @[simp] -lemma Sup_singleton {a : α} : Sup {a} = a := +theorem Sup_singleton {a : α} : Sup {a} = a := by finish [singleton_def] --eq.trans Sup_insert $ by simp @[simp] -lemma Inf_singleton {a : α} : Inf {a} = a := +theorem Inf_singleton {a : α} : Inf {a} = a := by finish [singleton_def] --eq.trans Inf_insert $ by simp @@ -167,52 +167,52 @@ variables [complete_lattice α] {s t : ι → α} {a b : α} -- TODO: this declaration gives error when starting smt state --@[ematch] -lemma le_supr (s : ι → α) (i : ι) : s i ≤ supr s := +theorem le_supr (s : ι → α) (i : ι) : s i ≤ supr s := le_Sup ⟨i, rfl⟩ @[ematch] -lemma le_supr' (s : ι → α) (i : ι) : (: s i ≤ supr s :) := +theorem le_supr' (s : ι → α) (i : ι) : (: s i ≤ supr s :) := le_Sup ⟨i, rfl⟩ /- TODO: this version would be more powerful, but, alas, the pattern matcher doesn't accept it. @[ematch] -lemma le_supr' (s : ι → α) (i : ι) : (: s i :) ≤ (: supr s :) := +theorem le_supr' (s : ι → α) (i : ι) : (: s i :) ≤ (: supr s :) := le_Sup ⟨i, rfl⟩ -/ -lemma le_supr_of_le (i : ι) (h : a ≤ s i) : a ≤ supr s := +theorem le_supr_of_le (i : ι) (h : a ≤ s i) : a ≤ supr s := le_trans h (le_supr _ i) -lemma supr_le (h : ∀i, s i ≤ a) : supr s ≤ a := +theorem supr_le (h : ∀i, s i ≤ a) : supr s ≤ a := Sup_le $ assume b ⟨i, eq⟩, eq.symm ▸ h i -lemma supr_le_supr (h : ∀i, s i ≤ t i) : supr s ≤ supr t := +theorem supr_le_supr (h : ∀i, s i ≤ t i) : supr s ≤ supr t := supr_le $ assume i, le_supr_of_le i (h i) -lemma supr_le_supr2 {t : ι₂ → α} (h : ∀i, ∃j, s i ≤ t j) : supr s ≤ supr t := +theorem supr_le_supr2 {t : ι₂ → α} (h : ∀i, ∃j, s i ≤ t j) : supr s ≤ supr t := supr_le $ assume j, exists.elim (h j) le_supr_of_le -lemma supr_le_supr_const (h : ι → ι₂) : (⨆ i:ι, a) ≤ (⨆ j:ι₂, a) := +theorem supr_le_supr_const (h : ι → ι₂) : (⨆ i:ι, a) ≤ (⨆ j:ι₂, a) := supr_le $ le_supr _ ∘ h @[simp] -lemma supr_le_iff : supr s ≤ a ↔ (∀i, s i ≤ a) := +theorem supr_le_iff : supr s ≤ a ↔ (∀i, s i ≤ a) := ⟨assume : supr s ≤ a, assume i, le_trans (le_supr _ _) this, supr_le⟩ -- TODO: finish doesn't do well here. @[congr] -lemma supr_congr_Prop {p q : Prop} {f₁ : p → α} {f₂ : q → α} +theorem supr_congr_Prop {p q : Prop} {f₁ : p → α} {f₂ : q → α} (pq : p ↔ q) (f : ∀x, f₁ (pq.mpr x) = f₂ x) : supr f₁ = supr f₂ := le_antisymm (supr_le_supr2 $ assume j, ⟨pq.mp j, le_of_eq $ f _⟩) (supr_le_supr2 $ assume j, ⟨pq.mpr j, le_of_eq $ (f j).symm⟩) -lemma infi_le (s : ι → α) (i : ι) : infi s ≤ s i := +theorem infi_le (s : ι → α) (i : ι) : infi s ≤ s i := Inf_le ⟨i, rfl⟩ @[ematch] -lemma infi_le' (s : ι → α) (i : ι) : (: infi s ≤ s i :) := +theorem infi_le' (s : ι → α) (i : ι) : (: infi s ≤ s i :) := Inf_le ⟨i, rfl⟩ example {f : β → α} (b : β) : (⨅ x, f x) ≤ f b := @@ -222,7 +222,7 @@ end /- I wanted to see if this would help for infi_comm; it doesn't. @[ematch] -lemma infi_le₂' (s : ι → ι₂ → α) (i : ι) (j : ι₂): (: ⨅ i j, s i j :) ≤ (: s i j :) := +theorem infi_le₂' (s : ι → ι₂ → α) (i : ι) (j : ι₂): (: ⨅ i j, s i j :) ≤ (: s i j :) := begin transitivity, apply (infi_le (λ i, ⨅ j, s i j) i), @@ -230,27 +230,27 @@ begin end -/ -lemma infi_le_of_le (i : ι) (h : s i ≤ a) : infi s ≤ a := +theorem infi_le_of_le (i : ι) (h : s i ≤ a) : infi s ≤ a := le_trans (infi_le _ i) h -lemma le_infi (h : ∀i, a ≤ s i) : a ≤ infi s := +theorem le_infi (h : ∀i, a ≤ s i) : a ≤ infi s := le_Inf $ assume b ⟨i, eq⟩, eq.symm ▸ h i -lemma infi_le_infi (h : ∀i, s i ≤ t i) : infi s ≤ infi t := +theorem infi_le_infi (h : ∀i, s i ≤ t i) : infi s ≤ infi t := le_infi $ assume i, infi_le_of_le i (h i) -lemma infi_le_infi2 {t : ι₂ → α} (h : ∀j, ∃i, s i ≤ t j) : infi s ≤ infi t := +theorem infi_le_infi2 {t : ι₂ → α} (h : ∀j, ∃i, s i ≤ t j) : infi s ≤ infi t := le_infi $ assume j, exists.elim (h j) infi_le_of_le -lemma infi_le_infi_const (h : ι₂ → ι) : (⨅ i:ι, a) ≤ (⨅ j:ι₂, a) := +theorem infi_le_infi_const (h : ι₂ → ι) : (⨅ i:ι, a) ≤ (⨅ j:ι₂, a) := le_infi $ infi_le _ ∘ h @[simp] -lemma le_infi_iff : a ≤ infi s ↔ (∀i, a ≤ s i) := +theorem le_infi_iff : a ≤ infi s ↔ (∀i, a ≤ s i) := ⟨assume : a ≤ infi s, assume i, le_trans this (infi_le _ _), le_infi⟩ @[congr] -lemma infi_congr_Prop {p q : Prop} {f₁ : p → α} {f₂ : q → α} +theorem infi_congr_Prop {p q : Prop} {f₁ : p → α} {f₂ : q → α} (pq : p ↔ q) (f : ∀x, f₁ (pq.mpr x) = f₂ x) : infi f₁ = infi f₂ := le_antisymm (infi_le_infi2 $ assume j, ⟨pq.mpr j, le_of_eq $ f j⟩) @@ -258,16 +258,16 @@ le_antisymm -- TODO: why isn't this a valid simp lemma? -- @[simp] -lemma infi_const {a : α} (b : ι) : (⨅ b:ι, a) = a := +theorem infi_const {a : α} (b : ι) : (⨅ b:ι, a) = a := le_antisymm (Inf_le ⟨b, rfl⟩) (by finish) -- le_antisymm (Inf_le ⟨b, rfl⟩) (le_Inf $ assume a' ⟨b', h⟩, h.symm ▸ le_refl _) -lemma supr_const {a : α} (b : ι) : (⨆ b:ι, a) = a := +theorem supr_const {a : α} (b : ι) : (⨆ b:ι, a) = a := le_antisymm (by finish) (le_Sup ⟨b, rfl⟩) --le_antisymm (Sup_le $ assume a' ⟨b', h⟩, h.symm ▸ le_refl _) (le_Sup ⟨b, rfl⟩) -- TODO: should this be @[simp]? -lemma infi_comm {f : ι → ι₂ → α} : (⨅i, ⨅j, f i j) = (⨅j, ⨅i, f i j) := +theorem infi_comm {f : ι → ι₂ → α} : (⨅i, ⨅j, f i j) = (⨅j, ⨅i, f i j) := le_antisymm (le_infi $ assume i, le_infi $ assume j, infi_le_of_le j $ infi_le _ i) (le_infi $ assume j, le_infi $ assume i, infi_le_of_le i $ infi_le _ j) @@ -285,31 +285,31 @@ end -/ -- TODO: should this be @[simp]? -lemma supr_comm {f : ι → ι₂ → α} : (⨆i, ⨆j, f i j) = (⨆j, ⨆i, f i j) := +theorem supr_comm {f : ι → ι₂ → α} : (⨆i, ⨆j, f i j) = (⨆j, ⨆i, f i j) := le_antisymm (supr_le $ assume i, supr_le $ assume j, le_supr_of_le j $ le_supr _ i) (supr_le $ assume j, supr_le $ assume i, le_supr_of_le i $ le_supr _ j) @[simp] -lemma infi_infi_eq_left {b : β} {f : Πx:β, x = b → α} : (⨅x, ⨅h:x = b, f x h) = f b rfl := +theorem infi_infi_eq_left {b : β} {f : Πx:β, x = b → α} : (⨅x, ⨅h:x = b, f x h) = f b rfl := le_antisymm (infi_le_of_le b $ infi_le _ rfl) (le_infi $ assume b', le_infi $ assume eq, match b', eq with ._, rfl := le_refl _ end) @[simp] -lemma infi_infi_eq_right {b : β} {f : Πx:β, b = x → α} : (⨅x, ⨅h:b = x, f x h) = f b rfl := +theorem infi_infi_eq_right {b : β} {f : Πx:β, b = x → α} : (⨅x, ⨅h:b = x, f x h) = f b rfl := le_antisymm (infi_le_of_le b $ infi_le _ rfl) (le_infi $ assume b', le_infi $ assume eq, match b', eq with ._, rfl := le_refl _ end) @[simp] -lemma supr_supr_eq_left {b : β} {f : Πx:β, x = b → α} : (⨆x, ⨆h : x = b, f x h) = f b rfl := +theorem supr_supr_eq_left {b : β} {f : Πx:β, x = b → α} : (⨆x, ⨆h : x = b, f x h) = f b rfl := le_antisymm (supr_le $ assume b', supr_le $ assume eq, match b', eq with ._, rfl := le_refl _ end) (le_supr_of_le b $ le_supr _ rfl) @[simp] -lemma supr_supr_eq_right {b : β} {f : Πx:β, b = x → α} : (⨆x, ⨆h : b = x, f x h) = f b rfl := +theorem supr_supr_eq_right {b : β} {f : Πx:β, b = x → α} : (⨆x, ⨆h : b = x, f x h) = f b rfl := le_antisymm (supr_le $ assume b', supr_le $ assume eq, match b', eq with ._, rfl := le_refl _ end) (le_supr_of_le b $ le_supr _ rfl) @@ -317,14 +317,14 @@ le_antisymm attribute [ematch] le_refl @[ematch] -lemma foo {a b : α} (h : a = b) : a ≤ b := +theorem foo {a b : α} (h : a = b) : a ≤ b := by rw h; apply le_refl @[ematch] -lemma foo' {a b : α} (h : b = a) : a ≤ b := +theorem foo' {a b : α} (h : b = a) : a ≤ b := by rw h; apply le_refl -lemma infi_inf_eq {f g : β → α} : (⨅ x, f x ⊓ g x) = (⨅ x, f x) ⊓ (⨅ x, g x) := +theorem infi_inf_eq {f g : β → α} : (⨅ x, f x ⊓ g x) = (⨅ x, f x) ⊓ (⨅ x, g x) := le_antisymm (le_inf (le_infi $ assume i, infi_le_of_le i inf_le_left) @@ -342,7 +342,7 @@ begin end -/ -lemma supr_sup_eq {f g : β → α} : (⨆ x, f x ⊔ g x) = (⨆ x, f x) ⊔ (⨆ x, g x) := +theorem supr_sup_eq {f g : β → α} : (⨆ x, f x ⊔ g x) = (⨆ x, f x) ⊔ (⨆ x, g x) := le_antisymm (supr_le $ assume i, sup_le (le_sup_left_of_le $ le_supr _ _) @@ -354,44 +354,44 @@ le_antisymm /- supr and infi under Prop -/ @[simp] -lemma infi_false {s : false → α} : infi s = ⊤ := +theorem infi_false {s : false → α} : infi s = ⊤ := le_antisymm le_top (le_infi $ assume i, false.elim i) @[simp] -lemma supr_false {s : false → α} : supr s = ⊥ := +theorem supr_false {s : false → α} : supr s = ⊥ := le_antisymm (supr_le $ assume i, false.elim i) bot_le @[simp] -lemma infi_true {s : true → α} : infi s = s trivial := +theorem infi_true {s : true → α} : infi s = s trivial := le_antisymm (infi_le _ _) (le_infi $ assume ⟨⟩, le_refl _) @[simp] -lemma supr_true {s : true → α} : supr s = s trivial := +theorem supr_true {s : true → α} : supr s = s trivial := le_antisymm (supr_le $ assume ⟨⟩, le_refl _) (le_supr _ _) @[simp] -lemma infi_exists {p : ι → Prop} {f : Exists p → α} : (⨅ x, f x) = (⨅ i, ⨅ h:p i, f ⟨i, h⟩) := +theorem infi_exists {p : ι → Prop} {f : Exists p → α} : (⨅ x, f x) = (⨅ i, ⨅ h:p i, f ⟨i, h⟩) := le_antisymm (le_infi $ assume i, le_infi $ assume : p i, infi_le _ _) (le_infi $ assume ⟨i, h⟩, infi_le_of_le i $ infi_le _ _) @[simp] -lemma supr_exists {p : ι → Prop} {f : Exists p → α} : (⨆ x, f x) = (⨆ i, ⨆ h:p i, f ⟨i, h⟩) := +theorem supr_exists {p : ι → Prop} {f : Exists p → α} : (⨆ x, f x) = (⨆ i, ⨆ h:p i, f ⟨i, h⟩) := le_antisymm (supr_le $ assume ⟨i, h⟩, le_supr_of_le i $ le_supr (λh:p i, f ⟨i, h⟩) _) (supr_le $ assume i, supr_le $ assume : p i, le_supr _ _) -lemma infi_and {p q : Prop} {s : p ∧ q → α} : infi s = (⨅ h₁ : p, ⨅ h₂ : q, s ⟨h₁, h₂⟩) := +theorem infi_and {p q : Prop} {s : p ∧ q → α} : infi s = (⨅ h₁ : p, ⨅ h₂ : q, s ⟨h₁, h₂⟩) := le_antisymm (le_infi $ assume i, le_infi $ assume j, infi_le _ _) (le_infi $ assume ⟨i, h⟩, infi_le_of_le i $ infi_le _ _) -lemma supr_and {p q : Prop} {s : p ∧ q → α} : supr s = (⨆ h₁ : p, ⨆ h₂ : q, s ⟨h₁, h₂⟩) := +theorem supr_and {p q : Prop} {s : p ∧ q → α} : supr s = (⨆ h₁ : p, ⨆ h₂ : q, s ⟨h₁, h₂⟩) := le_antisymm (supr_le $ assume ⟨i, h⟩, le_supr_of_le i $ le_supr (λj, s ⟨i, j⟩) _) (supr_le $ assume i, supr_le $ assume j, le_supr _ _) -lemma infi_or {p q : Prop} {s : p ∨ q → α} : +theorem infi_or {p q : Prop} {s : p ∨ q → α} : infi s = (⨅ h : p, s (or.inl h)) ⊓ (⨅ h : q, s (or.inr h)) := le_antisymm (le_inf @@ -402,7 +402,7 @@ le_antisymm | or.inr j := inf_le_right_of_le $ infi_le _ _ end) -lemma supr_or {p q : Prop} {s : p ∨ q → α} : +theorem supr_or {p q : Prop} {s : p ∨ q → α} : (⨆ x, s x) = (⨆ i, s (or.inl i)) ⊔ (⨆ j, s (or.inr j)) := le_antisymm (supr_le $ assume s, match s with @@ -413,24 +413,24 @@ le_antisymm (supr_le_supr2 $ assume i, ⟨or.inl i, le_refl _⟩) (supr_le_supr2 $ assume j, ⟨or.inr j, le_refl _⟩)) -lemma Inf_eq_infi {s : set α} : Inf s = (⨅a ∈ s, a) := +theorem Inf_eq_infi {s : set α} : Inf s = (⨅a ∈ s, a) := le_antisymm (le_infi $ assume b, le_infi $ assume h, Inf_le h) (le_Inf $ assume b h, infi_le_of_le b $ infi_le _ h) -lemma Sup_eq_supr {s : set α} : Sup s = (⨆a ∈ s, a) := +theorem Sup_eq_supr {s : set α} : Sup s = (⨆a ∈ s, a) := le_antisymm (Sup_le $ assume b h, le_supr_of_le b $ le_supr _ h) (supr_le $ assume b, supr_le $ assume h, le_Sup h) -lemma Inf_image {s : set β} {f : β → α} : Inf (set.image f s) = (⨅ a ∈ s, f a) := +theorem Inf_image {s : set β} {f : β → α} : Inf (set.image f s) = (⨅ a ∈ s, f a) := calc Inf (set.image f s) = (⨅a, ⨅h : ∃b, b ∈ s ∧ f b = a, a) : Inf_eq_infi ... = (⨅a, ⨅b, ⨅h : f b = a ∧ b ∈ s, a) : by simp ... = (⨅a, ⨅b, ⨅h : a = f b, ⨅h : b ∈ s, a) : by simp [infi_and, eq_comm] ... = (⨅b, ⨅a, ⨅h : a = f b, ⨅h : b ∈ s, a) : by rw [infi_comm] ... = (⨅a∈s, f a) : congr_arg infi $ funext $ assume x, by rw [infi_infi_eq_left] -lemma Sup_image {s : set β} {f : β → α} : Sup (set.image f s) = (⨆ a ∈ s, f a) := +theorem Sup_image {s : set β} {f : β → α} : Sup (set.image f s) = (⨆ a ∈ s, f a) := calc Sup (set.image f s) = (⨆a, ⨆h : ∃b, b ∈ s ∧ f b = a, a) : Sup_eq_supr ... = (⨆a, ⨆b, ⨆h : f b = a ∧ b ∈ s, a) : by simp ... = (⨆a, ⨆b, ⨆h : a = f b, ⨆h : b ∈ s, a) : by simp [supr_and, eq_comm] @@ -441,102 +441,102 @@ calc Sup (set.image f s) = (⨆a, ⨆h : ∃b, b ∈ s ∧ f b = a, a) : Sup_eq_ /- should work using the simplifier! -/ @[simp] -lemma infi_emptyset {f : β → α} : (⨅ x ∈ (∅ : set β), f x) = ⊤ := +theorem infi_emptyset {f : β → α} : (⨅ x ∈ (∅ : set β), f x) = ⊤ := le_antisymm le_top (le_infi $ assume x, le_infi false.elim) @[simp] -lemma supr_emptyset {f : β → α} : (⨆ x ∈ (∅ : set β), f x) = ⊥ := +theorem supr_emptyset {f : β → α} : (⨆ x ∈ (∅ : set β), f x) = ⊥ := le_antisymm (supr_le $ assume x, supr_le false.elim) bot_le @[simp] -lemma infi_univ {f : β → α} : (⨅ x ∈ (univ : set β), f x) = (⨅ x, f x) := +theorem infi_univ {f : β → α} : (⨅ x ∈ (univ : set β), f x) = (⨅ x, f x) := show (⨅ (x : β) (H : true), f x) = ⨅ (x : β), f x, from congr_arg infi $ funext $ assume x, infi_const ⟨⟩ @[simp] -lemma supr_univ {f : β → α} : (⨆ x ∈ (univ : set β), f x) = (⨆ x, f x) := +theorem supr_univ {f : β → α} : (⨆ x ∈ (univ : set β), f x) = (⨆ x, f x) := show (⨆ (x : β) (H : true), f x) = ⨆ (x : β), f x, from congr_arg supr $ funext $ assume x, supr_const ⟨⟩ @[simp] -lemma infi_union {f : β → α} {s t : set β} : (⨅ x ∈ s ∪ t, f x) = (⨅x∈s, f x) ⊓ (⨅x∈t, f x) := +theorem infi_union {f : β → α} {s t : set β} : (⨅ x ∈ s ∪ t, f x) = (⨅x∈s, f x) ⊓ (⨅x∈t, f x) := calc (⨅ x ∈ s ∪ t, f x) = (⨅ x, (⨅h : x∈s, f x) ⊓ (⨅h : x∈t, f x)) : congr_arg infi $ funext $ assume x, infi_or ... = (⨅x∈s, f x) ⊓ (⨅x∈t, f x) : infi_inf_eq @[simp] -lemma supr_union {f : β → α} {s t : set β} : (⨆ x ∈ s ∪ t, f x) = (⨆x∈s, f x) ⊔ (⨆x∈t, f x) := +theorem supr_union {f : β → α} {s t : set β} : (⨆ x ∈ s ∪ t, f x) = (⨆x∈s, f x) ⊔ (⨆x∈t, f x) := calc (⨆ x ∈ s ∪ t, f x) = (⨆ x, (⨆h : x∈s, f x) ⊔ (⨆h : x∈t, f x)) : congr_arg supr $ funext $ assume x, supr_or ... = (⨆x∈s, f x) ⊔ (⨆x∈t, f x) : supr_sup_eq @[simp] theorem insert_of_has_insert (x : α) (a : set α) : has_insert.insert x a = insert x a := rfl @[simp] -lemma infi_insert {f : β → α} {s : set β} {b : β} : (⨅ x ∈ insert b s, f x) = f b ⊓ (⨅x∈s, f x) := +theorem infi_insert {f : β → α} {s : set β} {b : β} : (⨅ x ∈ insert b s, f x) = f b ⊓ (⨅x∈s, f x) := eq.trans infi_union $ congr_arg (λx:α, x ⊓ (⨅x∈s, f x)) infi_infi_eq_left @[simp] -lemma supr_insert {f : β → α} {s : set β} {b : β} : (⨆ x ∈ insert b s, f x) = f b ⊔ (⨆x∈s, f x) := +theorem supr_insert {f : β → α} {s : set β} {b : β} : (⨆ x ∈ insert b s, f x) = f b ⊔ (⨆x∈s, f x) := eq.trans supr_union $ congr_arg (λx:α, x ⊔ (⨆x∈s, f x)) supr_supr_eq_left @[simp] -lemma infi_singleton {f : β → α} {b : β} : (⨅ x ∈ (singleton b : set β), f x) = f b := +theorem infi_singleton {f : β → α} {b : β} : (⨅ x ∈ (singleton b : set β), f x) = f b := show (⨅ x ∈ insert b (∅ : set β), f x) = f b, by simp @[simp] -lemma supr_singleton {f : β → α} {b : β} : (⨆ x ∈ (singleton b : set β), f x) = f b := +theorem supr_singleton {f : β → α} {b : β} : (⨆ x ∈ (singleton b : set β), f x) = f b := show (⨆ x ∈ insert b (∅ : set β), f x) = f b, by simp /- supr and infi under Type -/ @[simp] -lemma infi_empty {s : empty → α} : infi s = ⊤ := +theorem infi_empty {s : empty → α} : infi s = ⊤ := le_antisymm le_top (le_infi $ assume i, empty.rec_on _ i) @[simp] -lemma supr_empty {s : empty → α} : supr s = ⊥ := +theorem supr_empty {s : empty → α} : supr s = ⊥ := le_antisymm (supr_le $ assume i, empty.rec_on _ i) bot_le @[simp] -lemma infi_unit {f : unit → α} : (⨅ x, f x) = f () := +theorem infi_unit {f : unit → α} : (⨅ x, f x) = f () := le_antisymm (infi_le _ _) (le_infi $ assume ⟨⟩, le_refl _) @[simp] -lemma supr_unit {f : unit → α} : (⨆ x, f x) = f () := +theorem supr_unit {f : unit → α} : (⨆ x, f x) = f () := le_antisymm (supr_le $ assume ⟨⟩, le_refl _) (le_supr _ _) -lemma infi_subtype {p : ι → Prop} {f : subtype p → α} : (⨅ x, f x) = (⨅ i, ⨅ h:p i, f ⟨i, h⟩) := +theorem infi_subtype {p : ι → Prop} {f : subtype p → α} : (⨅ x, f x) = (⨅ i, ⨅ h:p i, f ⟨i, h⟩) := le_antisymm (le_infi $ assume i, le_infi $ assume : p i, infi_le _ _) (le_infi $ assume ⟨i, h⟩, infi_le_of_le i $ infi_le _ _) -lemma supr_subtype {p : ι → Prop} {f : subtype p → α} : (⨆ x, f x) = (⨆ i, ⨆ h:p i, f ⟨i, h⟩) := +theorem supr_subtype {p : ι → Prop} {f : subtype p → α} : (⨆ x, f x) = (⨆ i, ⨆ h:p i, f ⟨i, h⟩) := le_antisymm (supr_le $ assume ⟨i, h⟩, le_supr_of_le i $ le_supr (λh:p i, f ⟨i, h⟩) _) (supr_le $ assume i, supr_le $ assume : p i, le_supr _ _) -lemma infi_sigma {p : β → Type w} {f : sigma p → α} : (⨅ x, f x) = (⨅ i, ⨅ h:p i, f ⟨i, h⟩) := +theorem infi_sigma {p : β → Type w} {f : sigma p → α} : (⨅ x, f x) = (⨅ i, ⨅ h:p i, f ⟨i, h⟩) := le_antisymm (le_infi $ assume i, le_infi $ assume : p i, infi_le _ _) (le_infi $ assume ⟨i, h⟩, infi_le_of_le i $ infi_le _ _) -lemma supr_sigma {p : β → Type w} {f : sigma p → α} : (⨆ x, f x) = (⨆ i, ⨆ h:p i, f ⟨i, h⟩) := +theorem supr_sigma {p : β → Type w} {f : sigma p → α} : (⨆ x, f x) = (⨆ i, ⨆ h:p i, f ⟨i, h⟩) := le_antisymm (supr_le $ assume ⟨i, h⟩, le_supr_of_le i $ le_supr (λh:p i, f ⟨i, h⟩) _) (supr_le $ assume i, supr_le $ assume : p i, le_supr _ _) -lemma infi_prod {γ : Type w} {f : β × γ → α} : (⨅ x, f x) = (⨅ i, ⨅ j, f (i, j)) := +theorem infi_prod {γ : Type w} {f : β × γ → α} : (⨅ x, f x) = (⨅ i, ⨅ j, f (i, j)) := le_antisymm (le_infi $ assume i, le_infi $ assume j, infi_le _ _) (le_infi $ assume ⟨i, h⟩, infi_le_of_le i $ infi_le _ _) -lemma supr_prod {γ : Type w} {f : β × γ → α} : (⨆ x, f x) = (⨆ i, ⨆ j, f (i, j)) := +theorem supr_prod {γ : Type w} {f : β × γ → α} : (⨆ x, f x) = (⨆ i, ⨆ j, f (i, j)) := le_antisymm (supr_le $ assume ⟨i, h⟩, le_supr_of_le i $ le_supr (λj, f ⟨i, j⟩) _) (supr_le $ assume i, supr_le $ assume j, le_supr _ _) -lemma infi_sum {γ : Type w} {f : β ⊕ γ → α} : +theorem infi_sum {γ : Type w} {f : β ⊕ γ → α} : (⨅ x, f x) = (⨅ i, f (sum.inl i)) ⊓ (⨅ j, f (sum.inr j)) := le_antisymm (le_inf @@ -547,7 +547,7 @@ le_antisymm | sum.inr j := inf_le_right_of_le $ infi_le _ _ end) -lemma supr_sum {γ : Type w} {f : β ⊕ γ → α} : +theorem supr_sum {γ : Type w} {f : β ⊕ γ → α} : (⨆ x, f x) = (⨆ i, f (sum.inl i)) ⊔ (⨆ j, f (sum.inr j)) := le_antisymm (supr_le $ assume s, match s with @@ -584,10 +584,10 @@ instance complete_lattice_fun {α : Type u} {β : Type v} [complete_lattice β] section complete_lattice variables [weak_order α] [complete_lattice β] -lemma monotone_Sup_of_monotone {s : set (α → β)} (m_s : ∀f∈s, monotone f) : monotone (Sup s) := +theorem monotone_Sup_of_monotone {s : set (α → β)} (m_s : ∀f∈s, monotone f) : monotone (Sup s) := assume x y h, Sup_le $ assume x' ⟨f, f_in, fx_eq⟩, le_Sup_of_le ⟨f, f_in, rfl⟩ $ fx_eq ▸ m_s _ f_in h -lemma monotone_Inf_of_monotone {s : set (α → β)} (m_s : ∀f∈s, monotone f) : monotone (Inf s) := +theorem monotone_Inf_of_monotone {s : set (α → β)} (m_s : ∀f∈s, monotone f) : monotone (Inf s) := assume x y h, le_Inf $ assume x' ⟨f, f_in, fx_eq⟩, Inf_le_of_le ⟨f, f_in, rfl⟩ $ fx_eq ▸ m_s _ f_in h end complete_lattice @@ -598,19 +598,19 @@ end lattice /- Classical statements: @[simp] -lemma Inf_eq_top : Inf s = ⊤ ↔ (∀a∈s, a = ⊤) := +theorem Inf_eq_top : Inf s = ⊤ ↔ (∀a∈s, a = ⊤) := _ @[simp] -lemma infi_eq_top : infi s = ⊤ ↔ (∀i, s i = ⊤) := +theorem infi_eq_top : infi s = ⊤ ↔ (∀i, s i = ⊤) := _ @[simp] -lemma Sup_eq_bot : Sup s = ⊤ ↔ (∀a∈s, a = ⊥) := +theorem Sup_eq_bot : Sup s = ⊤ ↔ (∀a∈s, a = ⊥) := _ @[simp] -lemma supr_eq_top : supr s = ⊤ ↔ (∀i, s i = ⊥) := +theorem supr_eq_top : supr s = ⊤ ↔ (∀i, s i = ⊥) := _ diff --git a/algebra/lattice/filter.lean b/algebra/lattice/filter.lean index 2b61cc2122ba6..2e30e2f97fcad 100644 --- a/algebra/lattice/filter.lean +++ b/algebra/lattice/filter.lean @@ -13,7 +13,7 @@ universes u v w x section applicative variables {f : Type u → Type v} [applicative f] {α β : Type u} -lemma pure_seq_eq_map : ∀ {α β : Type u} (g : α → β) (x : f α), pure g <*> x = g <$> x := +theorem pure_seq_eq_map : ∀ {α β : Type u} (g : α → β) (x : f α), pure g <*> x = g <$> x := @applicative.pure_seq_eq_map f _ end applicative @@ -41,26 +41,26 @@ section prod variables {α : Type u} {β : Type v} @[simp] -- copied from parser -lemma prod.mk.eta : ∀{p : α × β}, (p.1, p.2) = p +theorem prod.mk.eta : ∀{p : α × β}, (p.1, p.2) = p | (a, b) := rfl def prod.swap : (α×β) → (β×α) := λp, (p.2, p.1) @[simp] -lemma prod.swap_swap : ∀x:α×β, prod.swap (prod.swap x) = x +theorem prod.swap_swap : ∀x:α×β, prod.swap (prod.swap x) = x | ⟨a, b⟩ := rfl @[simp] -lemma prod.fst_swap {p : α×β} : (prod.swap p).1 = p.2 := rfl +theorem prod.fst_swap {p : α×β} : (prod.swap p).1 = p.2 := rfl @[simp] -lemma prod.snd_swap {p : α×β} : (prod.swap p).2 = p.1 := rfl +theorem prod.snd_swap {p : α×β} : (prod.swap p).2 = p.1 := rfl @[simp] -lemma prod.swap_prod_mk {a : α} {b : β} : prod.swap (a, b) = (b, a) := rfl +theorem prod.swap_prod_mk {a : α} {b : β} : prod.swap (a, b) = (b, a) := rfl @[simp] -lemma prod.swap_swap_eq : prod.swap ∘ prod.swap = @id (α × β) := +theorem prod.swap_swap_eq : prod.swap ∘ prod.swap = @id (α × β) := funext $ prod.swap_swap end prod @@ -68,13 +68,13 @@ end prod namespace lattice variables {α : Type u} {ι : Sort v} [complete_lattice α] -lemma Inf_eq_finite_sets {s : set α} : +theorem Inf_eq_finite_sets {s : set α} : Inf s = (⨅ t ∈ { t | finite t ∧ t ⊆ s}, Inf t) := le_antisymm (le_infi $ assume t, le_infi $ assume ⟨_, h⟩, Inf_le_Inf h) (le_Inf $ assume b h, infi_le_of_le {b} $ infi_le_of_le (by simp [h]) $ Inf_le $ by simp) -lemma Sup_le_iff {s : set α} {a : α} : Sup s ≤ a ↔ (∀x∈s, x ≤ a) := +theorem Sup_le_iff {s : set α} {a : α} : Sup s ≤ a ↔ (∀x∈s, x ≤ a) := ⟨assume h x hx, le_trans (le_Sup hx) h, Sup_le⟩ end lattice @@ -101,10 +101,10 @@ variables {α β : Type u} theorem ne_empty_iff_exists_mem {s : set α} : s ≠ ∅ ↔ ∃ x, x ∈ s := ⟨exists_mem_of_ne_empty, assume ⟨x, (hx : x ∈ s)⟩ h, by rw [h] at hx; assumption⟩ -lemma fmap_eq_image {f : α → β} {s : set α} : f <$> s = f '' s := +theorem fmap_eq_image {f : α → β} {s : set α} : f <$> s = f '' s := rfl -lemma mem_seq_iff {f : set (α → β)} {s : set α} {b : β} : +theorem mem_seq_iff {f : set (α → β)} {s : set α} {b : β} : b ∈ (f <*> s) ↔ (∃(f' : α → β), ∃a ∈ s, f' ∈ f ∧ b = f' a) := begin simp [seq_eq_bind_map], @@ -121,19 +121,19 @@ variables {α : Type u} {β : Type v} {γ : Type w} {δ : Type x} protected def prod (s : set α) (t : set β) : set (α × β) := {p | p.1 ∈ s ∧ p.2 ∈ t} -lemma mem_prod_eq {s : set α} {t : set β} {p : α × β} : +theorem mem_prod_eq {s : set α} {t : set β} {p : α × β} : p ∈ set.prod s t = (p.1 ∈ s ∧ p.2 ∈ t) := rfl -lemma prod_vimage_eq {s : set α} {t : set β} {f : γ → α} {g : δ → β} : +theorem prod_vimage_eq {s : set α} {t : set β} {f : γ → α} {g : δ → β} : set.prod (vimage f s) (vimage g t) = vimage (λp, (f p.1, g p.2)) (set.prod s t) := rfl -lemma prod_mono {s₁ s₂ : set α} {t₁ t₂ : set β} (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : +theorem prod_mono {s₁ s₂ : set α} {t₁ t₂ : set β} (hs : s₁ ⊆ s₂) (ht : t₁ ⊆ t₂) : set.prod s₁ t₁ ⊆ set.prod s₂ t₂ := assume x ⟨h₁, h₂⟩, ⟨hs h₁, ht h₂⟩ -lemma prod_inter_prod {s₁ s₂ : set α} {t₁ t₂ : set β} : +theorem prod_inter_prod {s₁ s₂ : set α} {t₁ t₂ : set β} : set.prod s₁ t₁ ∩ set.prod s₂ t₂ = set.prod (s₁ ∩ s₂) (t₁ ∩ t₂) := subset.antisymm (assume ⟨a, b⟩ ⟨⟨ha₁, hb₁⟩, ⟨ha₂, hb₂⟩⟩, ⟨⟨ha₁, ha₂⟩, ⟨hb₁, hb₂⟩⟩) @@ -141,17 +141,17 @@ subset.antisymm (prod_mono (inter_subset_left _ _) (inter_subset_left _ _)) (prod_mono (inter_subset_right _ _) (inter_subset_right _ _))) -lemma monotone_prod [weak_order α] {f : α → set β} {g : α → set γ} +theorem monotone_prod [weak_order α] {f : α → set β} {g : α → set γ} (hf : monotone f) (hg : monotone g) : monotone (λx, set.prod (f x) (g x)) := assume a b h, prod_mono (hf h) (hg h) -lemma image_swap_prod {s : set α} {t : set β} : +theorem image_swap_prod {s : set α} {t : set β} : image (λp:β×α, (p.2, p.1)) (set.prod t s) = set.prod s t := set.ext $ assume ⟨a, b⟩, by simp [mem_image_eq, set.prod]; exact ⟨ assume ⟨b', a', h_a, h_b, h⟩, by rw [h_a, h_b] at h; assumption, assume ⟨ha, hb⟩, ⟨b, a, rfl, rfl, ⟨ha, hb⟩⟩⟩ -lemma prod_image_image_eq {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} +theorem prod_image_image_eq {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} {s₁ : set α₁} {s₂ : set α₂} {m₁ : α₁ → β₁} {m₂ : α₂ → β₂} : set.prod (image m₁ s₁) (image m₂ s₂) = image (λp:α₁×α₂, (m₁ p.1, m₂ p.2)) (set.prod s₁ s₂) := set.ext $ assume ⟨b₁, b₂⟩, @@ -162,11 +162,11 @@ set.ext $ assume ⟨b₁, b₂⟩, assume ⟨⟨a₁, a₂⟩, ⟨ha₁, ha₂⟩, eq⟩, eq ▸ ⟨mem_image_of_mem m₁ ha₁, mem_image_of_mem m₂ ha₂⟩⟩ @[simp] -lemma prod_singleton_singleton {a : α} {b : β} : +theorem prod_singleton_singleton {a : α} {b : β} : set.prod {a} {b} = ({(a, b)} : set (α×β)) := set.ext $ assume ⟨a', b'⟩, by simp [set.prod] -lemma prod_neq_empty_iff {s : set α} {t : set β} : +theorem prod_neq_empty_iff {s : set α} {t : set β} : set.prod s t ≠ ∅ ↔ (s ≠ ∅ ∧ t ≠ ∅) := begin rw [ne_empty_iff_exists_mem, ne_empty_iff_exists_mem, ne_empty_iff_exists_mem, @@ -176,42 +176,42 @@ begin end @[simp] -lemma prod_mk_mem_set_prod_eq {a : α} {b : β} {s : set α} {t : set β} : +theorem prod_mk_mem_set_prod_eq {a : α} {b : β} {s : set α} {t : set β} : (a, b) ∈ set.prod s t = (a ∈ s ∧ b ∈ t) := rfl -lemma monotone_inter [weak_order β] {f g : β → set α} +theorem monotone_inter [weak_order β] {f g : β → set α} (hf : monotone f) (hg : monotone g) : monotone (λx, (f x) ∩ (g x)) := assume a b h x ⟨h₁, h₂⟩, ⟨hf h h₁, hg h h₂⟩ @[simp] -lemma vimage_set_of_eq {p : α → Prop} {f : β → α} : +theorem vimage_set_of_eq {p : α → Prop} {f : β → α} : vimage f {a | p a} = {a | p (f a)} := rfl @[simp] -lemma set_of_mem_eq {s : set α} : {x | x ∈ s} = s := +theorem set_of_mem_eq {s : set α} : {x | x ∈ s} = s := rfl -lemma mem_image_iff_of_inverse (f : α → β) (g : β → α) {b : β} {s : set α} +theorem mem_image_iff_of_inverse (f : α → β) (g : β → α) {b : β} {s : set α} (h₁ : ∀a, g (f a) = a ) (h₂ : ∀b, f (g b) = b ) : b ∈ f '' s ↔ g b ∈ s := ⟨assume ⟨a, ha, fa_eq⟩, fa_eq ▸ (h₁ a).symm ▸ ha, assume h, ⟨g b, h, h₂ b⟩⟩ -lemma image_eq_vimage_of_inverse (f : α → β) (g : β → α) +theorem image_eq_vimage_of_inverse (f : α → β) (g : β → α) (h₁ : ∀a, g (f a) = a ) (h₂ : ∀b, f (g b) = b ) : image f = vimage g := funext $ assume s, set.ext $ assume b, mem_image_iff_of_inverse f g h₁ h₂ -lemma image_swap_eq_vimage_swap : image (@prod.swap α β) = vimage prod.swap := +theorem image_swap_eq_vimage_swap : image (@prod.swap α β) = vimage prod.swap := image_eq_vimage_of_inverse (@prod.swap α β) (@prod.swap β α) begin simp; intros; trivial end begin simp; intros; trivial end -lemma monotone_set_of [weak_order α] {p : α → β → Prop} +theorem monotone_set_of [weak_order α] {p : α → β → Prop} (hp : ∀b, monotone (λa, p a b)) : monotone (λa, {b | p a b}) := assume a a' h b, hp b h -lemma diff_right_antimono {s t u : set α} (h : t ⊆ u) : s - u ⊆ s - t := +theorem diff_right_antimono {s t u : set α} (h : t ⊆ u) : s - u ⊆ s - t := assume x ⟨hs, hnx⟩, ⟨hs, assume hx, hnx $ h hx⟩ end set @@ -219,45 +219,45 @@ end set section variables {α : Type u} {ι : Sort v} -lemma sUnion_mono {s t : set (set α)} (h : s ⊆ t) : (⋃₀ s) ⊆ (⋃₀ t) := +theorem sUnion_mono {s t : set (set α)} (h : s ⊆ t) : (⋃₀ s) ⊆ (⋃₀ t) := sUnion_subset $ assume t' ht', subset_sUnion_of_mem $ h ht' -lemma Union_subset_Union {s t : ι → set α} (h : ∀i, s i ⊆ t i) : (⋃i, s i) ⊆ (⋃i, t i) := +theorem Union_subset_Union {s t : ι → set α} (h : ∀i, s i ⊆ t i) : (⋃i, s i) ⊆ (⋃i, t i) := @supr_le_supr (set α) ι _ s t h -lemma Union_subset_Union2 {ι₂ : Sort x} {s : ι → set α} {t : ι₂ → set α} (h : ∀i, ∃j, s i ⊆ t j) : (⋃i, s i) ⊆ (⋃i, t i) := +theorem Union_subset_Union2 {ι₂ : Sort x} {s : ι → set α} {t : ι₂ → set α} (h : ∀i, ∃j, s i ⊆ t j) : (⋃i, s i) ⊆ (⋃i, t i) := @supr_le_supr2 (set α) ι ι₂ _ s t h -lemma Union_subset_Union_const {ι₂ : Sort x} {s : set α} (h : ι → ι₂) : (⋃ i:ι, s) ⊆ (⋃ j:ι₂, s) := +theorem Union_subset_Union_const {ι₂ : Sort x} {s : set α} (h : ι → ι₂) : (⋃ i:ι, s) ⊆ (⋃ j:ι₂, s) := @supr_le_supr_const (set α) ι ι₂ _ s h -lemma diff_neq_empty {s t : set α} : s - t = ∅ ↔ s ⊆ t := +theorem diff_neq_empty {s t : set α} : s - t = ∅ ↔ s ⊆ t := ⟨assume h x hx, classical.by_contradiction $ assume : x ∉ t, show x ∈ (∅ : set α), from h ▸ ⟨hx, this⟩, assume h, bot_unique $ assume x ⟨hx, hnx⟩, hnx $ h hx⟩ @[simp] -lemma diff_empty {s : set α} : s - ∅ = s := +theorem diff_empty {s : set α} : s - ∅ = s := set.ext $ assume x, ⟨assume ⟨hx, _⟩, hx, assume h, ⟨h, not_false⟩⟩ end @[simp] -- should be handled by implies_true_iff -lemma implies_implies_true_iff {α : Sort u} {β : Sort v} : (α → β → true) ↔ true := +theorem implies_implies_true_iff {α : Sort u} {β : Sort v} : (α → β → true) ↔ true := ⟨assume _, trivial, assume _ _ _ , trivial⟩ @[simp] -lemma not_not_mem_iff {α : Type u} {a : α} {s : set α} [decidable (a ∈ s)] : +theorem not_not_mem_iff {α : Type u} {a : α} {s : set α} [decidable (a ∈ s)] : ¬ (a ∉ s) ↔ a ∈ s := not_not_iff @[simp] -lemma singleton_neq_emptyset {α : Type u} {a : α} : {a} ≠ (∅ : set α) := +theorem singleton_neq_emptyset {α : Type u} {a : α} : {a} ≠ (∅ : set α) := assume h, have a ∉ ({a} : set α), by simp [h], this $ mem_singleton a -lemma eq_of_sup_eq_inf_eq {α : Type u} [distrib_lattice α] {a b c : α} +theorem eq_of_sup_eq_inf_eq {α : Type u} [distrib_lattice α] {a b c : α} (h₁ : b ⊓ a = c ⊓ a) (h₂ : b ⊔ a = c ⊔ a) : b = c := le_antisymm (calc b ≤ (c ⊓ a) ⊔ b : le_sup_right @@ -269,7 +269,7 @@ le_antisymm ... = b ⊔ (b ⊓ a) : by rw [h₁, sup_inf_left, h₂]; simp [sup_comm] ... = b : sup_inf_self) -lemma inf_eq_bot_iff_le_compl {α : Type u} [bounded_distrib_lattice α] {a b c : α} +theorem inf_eq_bot_iff_le_compl {α : Type u} [bounded_distrib_lattice α] {a b c : α} (h₁ : b ⊔ c = ⊤) (h₂ : b ⊓ c = ⊥) : a ⊓ b = ⊥ ↔ a ≤ c := ⟨assume : a ⊓ b = ⊥, calc a ≤ a ⊓ (b ⊔ c) : by simp [h₁] @@ -280,16 +280,16 @@ lemma inf_eq_bot_iff_le_compl {α : Type u} [bounded_distrib_lattice α] {a b c calc a ⊓ b ≤ b ⊓ c : by rw [inf_comm]; exact inf_le_inf (le_refl _) this ... = ⊥ : h₂⟩ -lemma compl_image_set_of {α : Type u} {p : set α → Prop} : +theorem compl_image_set_of {α : Type u} {p : set α → Prop} : compl '' {x | p x} = {x | p (- x)} := set.ext $ assume x, ⟨assume ⟨y, (hy : p y), (h_eq : -y = x)⟩, show p (- x), by rw [←h_eq, lattice.neg_neg]; assumption, assume h : p (-x), ⟨_, h, lattice.neg_neg⟩⟩ -lemma neg_subset_neg_iff_subset {α : Type u} {x y : set α} : - y ⊆ - x ↔ x ⊆ y := +theorem neg_subset_neg_iff_subset {α : Type u} {x y : set α} : - y ⊆ - x ↔ x ⊆ y := @neg_le_neg_iff_le (set α) _ _ _ -lemma sUnion_eq_Union {α : Type u} {s : set (set α)} : +theorem sUnion_eq_Union {α : Type u} {s : set (set α)} : (⋃₀ s) = (⋃ (i : set α) (h : i ∈ s), i) := set.ext $ by simp @@ -300,7 +300,7 @@ local infix `≼` : 50 := r def directed {ι : Sort v} (f : ι → α) := ∀x, ∀y, ∃z, f z ≼ f x ∧ f z ≼ f y def directed_on (s : set α) := ∀x ∈ s, ∀y ∈ s, ∃z ∈ s, z ≼ x ∧ z ≼ y -lemma directed_on_Union {r} {ι : Sort v} {f : ι → set α} (hd : directed (⊇) f) +theorem directed_on_Union {r} {ι : Sort v} {f : ι → set α} (hd : directed (⊇) f) (h : ∀x, directed_on r (f x)) : directed_on r (⋃x, f x) := by simp [directed_on]; exact assume a₁ ⟨b₁, fb₁⟩ a₂ ⟨b₂, fb₂⟩, @@ -314,7 +314,7 @@ def upwards (s : set α) := ∀{x y}, x ∈ s → x ≼ y → y ∈ s end order -lemma directed_of_chain {α : Type u} {β : Type v} [weak_order β] {f : α → β} {c : set α} +theorem directed_of_chain {α : Type u} {β : Type v} [weak_order β] {f : α → β} {c : set α} (h : @zorn.chain α (λa b, f b ≤ f a) c) : directed (≤) (λx:{a:α // a ∈ c}, f (x.val)) := assume ⟨a, ha⟩ ⟨b, hb⟩, classical.by_cases @@ -334,22 +334,22 @@ structure filter (α : Type u) := namespace filter variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} -lemma filter_eq : ∀{f g : filter α}, f.sets = g.sets → f = g +theorem filter_eq : ∀{f g : filter α}, f.sets = g.sets → f = g | ⟨a, _, _, _⟩ ⟨._, _, _, _⟩ rfl := rfl -lemma univ_mem_sets' {f : filter α} {s : set α} (h : ∀ a, a ∈ s): s ∈ f.sets := +theorem univ_mem_sets' {f : filter α} {s : set α} (h : ∀ a, a ∈ s): s ∈ f.sets := let ⟨x, x_in_s⟩ := f.inhabited in f.upwards_sets x_in_s (assume x _, h x) -lemma univ_mem_sets {f : filter α} : univ ∈ f.sets := +theorem univ_mem_sets {f : filter α} : univ ∈ f.sets := univ_mem_sets' mem_univ -lemma inter_mem_sets {f : filter α} {x y : set α} (hx : x ∈ f.sets) (hy : y ∈ f.sets) : +theorem inter_mem_sets {f : filter α} {x y : set α} (hx : x ∈ f.sets) (hy : y ∈ f.sets) : x ∩ y ∈ f.sets := let ⟨z, ⟨z_in_s, z_le_x, z_le_y⟩⟩ := f.directed_sets _ hx _ hy in f.upwards_sets z_in_s (subset_inter z_le_x z_le_y) -lemma Inter_mem_sets {f : filter α} {s : β → set α} +theorem Inter_mem_sets {f : filter α} {s : β → set α} {is : set β} (hf : finite is) (hs : ∀i∈is, s i ∈ f.sets) : (⋂i∈is, s i) ∈ f.sets := begin /- equation compiler complains that this is requires well-founded recursion -/ induction hf with i is _ hf hi, @@ -363,12 +363,12 @@ begin /- equation compiler complains that this is requires well-founded recursio end end -lemma exists_sets_subset_iff {f : filter α} {x : set α} : +theorem exists_sets_subset_iff {f : filter α} {x : set α} : (∃y∈f.sets, y ⊆ x) ↔ x ∈ f.sets := ⟨assume ⟨y, hy, yx⟩, f.upwards_sets hy yx, assume hx, ⟨x, hx, subset.refl _⟩⟩ -lemma monotone_mem_sets {f : filter α} : monotone (λs, s ∈ f.sets) := +theorem monotone_mem_sets {f : filter α} : monotone (λs, s ∈ f.sets) := assume s t hst h, f.upwards_sets h hst def principal (s : set α) : filter α := @@ -451,33 +451,33 @@ instance : has_Sup (filter α) := ⟨join ∘ principal⟩ instance inhabited' : _root_.inhabited (filter α) := ⟨principal ∅⟩ -protected lemma le_Sup {s : set (filter α)} {f : filter α} : f ∈ s → f ≤ Sup s := +protected theorem le_Sup {s : set (filter α)} {f : filter α} : f ∈ s → f ≤ Sup s := assume f_in_s t' h, h f_in_s -protected lemma Sup_le {s : set (filter α)} {f : filter α} : (∀g∈s, g ≤ f) → Sup s ≤ f := +protected theorem Sup_le {s : set (filter α)} {f : filter α} : (∀g∈s, g ≤ f) → Sup s ≤ f := assume h a ha g hg, h g hg ha @[simp] -lemma mem_join_sets {s : set α} {f : filter (filter α)} : s ∈ (join f).sets = ({t | s ∈ filter.sets t} ∈ f.sets) := +theorem mem_join_sets {s : set α} {f : filter (filter α)} : s ∈ (join f).sets = ({t | s ∈ filter.sets t} ∈ f.sets) := rfl @[simp] -lemma mem_principal_sets {s t : set α} : s ∈ (principal t).sets = (t ⊆ s) := +theorem mem_principal_sets {s t : set α} : s ∈ (principal t).sets = (t ⊆ s) := rfl @[simp] -lemma le_principal_iff {s : set α} {f : filter α} : f ≤ principal s ↔ s ∈ f.sets := +theorem le_principal_iff {s : set α} {f : filter α} : f ≤ principal s ↔ s ∈ f.sets := show (∀{t}, s ⊆ t → t ∈ f.sets) ↔ s ∈ f.sets, from ⟨assume h, h (subset.refl s), assume hs t ht, f.upwards_sets hs ht⟩ -lemma principal_mono {s t : set α} : principal s ≤ principal t ↔ s ⊆ t := +theorem principal_mono {s t : set α} : principal s ≤ principal t ↔ s ⊆ t := by simp -lemma monotone_principal : monotone (principal : set α → filter α) := +theorem monotone_principal : monotone (principal : set α → filter α) := by simp [monotone, principal_mono]; exact assume a b h, h @[simp] -lemma principal_eq_iff_eq {s t : set α} : principal s = principal t ↔ s = t := +theorem principal_eq_iff_eq {s t : set α} : principal s = principal t ↔ s = t := by simp [eq_iff_le_and_le]; refl instance complete_lattice_filter : complete_lattice (filter α) := @@ -503,12 +503,12 @@ instance complete_lattice_filter : complete_lattice (filter α) := Inf_le := assume s a ha, filter.Sup_le $ assume b h, h _ ha } @[simp] -lemma map_principal {s : set α} {f : α → β} : +theorem map_principal {s : set α} {f : α → β} : map f (principal s) = principal (set.image f s) := filter_eq $ set.ext $ assume a, image_subset_iff_subset_vimage.symm @[simp] -lemma join_principal_eq_Sup {s : set (filter α)} : join (principal s) = Sup s := +theorem join_principal_eq_Sup {s : set (filter α)} : join (principal s) = Sup s := rfl instance monad_filter : monad filter := @@ -535,57 +535,57 @@ def at_bot [weak_order α] : filter α := ⨅ a, principal {b | b ≤ a} /- lattice equations -/ -lemma mem_inf_sets_of_left {f g : filter α} {s : set α} : +theorem mem_inf_sets_of_left {f g : filter α} {s : set α} : s ∈ f.sets → s ∈ (f ⊓ g).sets := have f ⊓ g ≤ f, from inf_le_left, assume hs, this hs -lemma mem_inf_sets_of_right {f g : filter α} {s : set α} : +theorem mem_inf_sets_of_right {f g : filter α} {s : set α} : s ∈ g.sets → s ∈ (f ⊓ g).sets := have f ⊓ g ≤ g, from inf_le_right, assume hs, this hs @[simp] -lemma mem_bot_sets {s : set α} : s ∈ (⊥ : filter α).sets := +theorem mem_bot_sets {s : set α} : s ∈ (⊥ : filter α).sets := assume x, false.elim -lemma empty_in_sets_eq_bot {f : filter α} : ∅ ∈ f.sets ↔ f = ⊥ := +theorem empty_in_sets_eq_bot {f : filter α} : ∅ ∈ f.sets ↔ f = ⊥ := ⟨assume h, bot_unique $ assume s _, f.upwards_sets h (empty_subset s), assume : f = ⊥, this.symm ▸ mem_bot_sets⟩ -lemma inhabited_of_mem_sets {f : filter α} {s : set α} (hf : f ≠ ⊥) (hs : s ∈ f.sets) : +theorem inhabited_of_mem_sets {f : filter α} {s : set α} (hf : f ≠ ⊥) (hs : s ∈ f.sets) : ∃x, x ∈ s := have ∅ ∉ f.sets, from assume h, hf $ empty_in_sets_eq_bot.mp h, have s ≠ ∅, from assume h, this (h ▸ hs), exists_mem_of_ne_empty this -lemma filter_eq_bot_of_not_nonempty {f : filter α} (ne : ¬ nonempty α) : f = ⊥ := +theorem filter_eq_bot_of_not_nonempty {f : filter α} (ne : ¬ nonempty α) : f = ⊥ := empty_in_sets_eq_bot.mp $ f.upwards_sets univ_mem_sets $ assume x, false.elim (ne ⟨x⟩) -lemma forall_sets_neq_empty_iff_neq_bot {f : filter α} : +theorem forall_sets_neq_empty_iff_neq_bot {f : filter α} : (∀ (s : set α), s ∈ f.sets → s ≠ ∅) ↔ f ≠ ⊥ := by simp [(@empty_in_sets_eq_bot α f).symm]; exact ⟨assume h hs, h _ hs rfl, assume h s hs eq, h $ eq ▸ hs⟩ -lemma mem_sets_of_neq_bot {f : filter α} {s : set α} (h : f ⊓ principal (-s) = ⊥) : s ∈ f.sets := +theorem mem_sets_of_neq_bot {f : filter α} {s : set α} (h : f ⊓ principal (-s) = ⊥) : s ∈ f.sets := have ∅ ∈ (f ⊓ principal (- s)).sets, from h.symm ▸ mem_bot_sets, let ⟨s₁, hs₁, s₂, (hs₂ : -s ⊆ s₂), (hs : s₁ ∩ s₂ ⊆ ∅)⟩ := this in have s₁ ⊆ s, from assume a ha, classical.by_contradiction $ assume ha', hs ⟨ha, hs₂ ha'⟩, f.upwards_sets hs₁ this @[simp] -lemma mem_sup_sets {f g : filter α} {s : set α} : +theorem mem_sup_sets {f g : filter α} {s : set α} : s ∈ (f ⊔ g).sets = (s ∈ f.sets ∧ s ∈ g.sets) := by refl @[simp] -lemma mem_inf_sets {f g : filter α} {s : set α} : +theorem mem_inf_sets {f g : filter α} {s : set α} : s ∈ (f ⊓ g).sets = (∃t₁∈f.sets, ∃t₂∈g.sets, t₁ ∩ t₂ ⊆ s) := by refl -lemma infi_sets_eq {f : ι → filter α} (h : directed (≤) f) (ne : nonempty ι) : +theorem infi_sets_eq {f : ι → filter α} (h : directed (≤) f) (ne : nonempty ι) : (infi f).sets = (⋃ i, (f i).sets) := let ⟨i⟩ := ne, @@ -599,7 +599,7 @@ in (show u ≤ infi f, from le_infi $ assume i, le_supr (λi, (f i).sets) i) (Union_subset $ assume i, infi_le f i) -lemma infi_sets_eq' {f : β → filter α} {s : set β} (h : directed_on (λx y, f x ≤ f y) s) (ne : ∃i, i ∈ s) : +theorem infi_sets_eq' {f : β → filter α} {s : set β} (h : directed_on (λx y, f x ≤ f y) s) (ne : ∃i, i ∈ s) : (⨅ i∈s, f i).sets = (⋃ i ∈ s, (f i).sets) := let ⟨i, hi⟩ := ne in calc (⨅ i ∈ s, f i).sets = (⨅ t : {t // t ∈ s}, (f t.val)).sets : by simp [infi_subtype]; refl @@ -608,7 +608,7 @@ calc (⨅ i ∈ s, f i).sets = (⨅ t : {t // t ∈ s}, (f t.val)).sets : by si ⟨⟨i, hi⟩⟩ ... = (⨆ t ∈ {t | t ∈ s}, (f t).sets) : by simp [supr_subtype]; refl -lemma Inf_sets_eq_finite {s : set (filter α)} : +theorem Inf_sets_eq_finite {s : set (filter α)} : (complete_lattice.Inf s).sets = (⋃ t ∈ {t | finite t ∧ t ⊆ s}, (Inf t).sets) := calc (Inf s).sets = (⨅ t ∈ { t | finite t ∧ t ⊆ s}, Inf t).sets : by rw [lattice.Inf_eq_finite_sets] ... = (⨆ t ∈ {t | finite t ∧ t ⊆ s}, (Inf t).sets) : infi_sets_eq' @@ -616,7 +616,7 @@ calc (Inf s).sets = (⨅ t ∈ { t | finite t ∧ t ⊆ s}, Inf t).sets : by rw Inf_le_Inf $ subset_union_left _ _, Inf_le_Inf $ subset_union_right _ _⟩) ⟨∅, by simp⟩ -lemma supr_sets_eq {f : ι → filter α} : (supr f).sets = (⋂i, (f i).sets) := +theorem supr_sets_eq {f : ι → filter α} : (supr f).sets = (⋂i, (f i).sets) := set.ext $ assume s, show s ∈ (join (principal {a : filter α | ∃i : ι, a = f i})).sets ↔ s ∈ (⋂i, (f i).sets), begin @@ -626,11 +626,11 @@ begin end @[simp] -lemma sup_join {f₁ f₂ : filter (filter α)} : (join f₁ ⊔ join f₂) = join (f₁ ⊔ f₂) := +theorem sup_join {f₁ f₂ : filter (filter α)} : (join f₁ ⊔ join f₂) = join (f₁ ⊔ f₂) := filter_eq $ set.ext $ assume x, by simp [supr_sets_eq, join] @[simp] -lemma supr_join {ι : Sort w} {f : ι → filter (filter α)} : (⨆x, join (f x)) = join (⨆x, f x) := +theorem supr_join {ι : Sort w} {f : ι → filter (filter α)} : (⨆x, join (f x)) = join (⨆x, f x) := filter_eq $ set.ext $ assume x, by simp [supr_sets_eq, join] instance : bounded_distrib_lattice (filter α) := @@ -648,7 +648,7 @@ instance : bounded_distrib_lattice (filter α) := subset.trans (@le_sup_inf (set α) _ _ _ _) (union_subset (subset.refl _) hs)⟩ end } -private theorem infi_finite_distrib {s : set (filter α)} {f : filter α} (h : finite s) : +private lemma infi_finite_distrib {s : set (filter α)} {f : filter α} (h : finite s) : (⨅ a ∈ s, f ⊔ a) = f ⊔ (Inf s) := begin induction h with a s hn hs hi, @@ -657,7 +657,7 @@ begin end /- the complementary version with ⨆ g∈s, f ⊓ g does not hold! -/ -lemma binfi_sup_eq { f : filter α } {s : set (filter α)} : +theorem binfi_sup_eq { f : filter α } {s : set (filter α)} : (⨅ g∈s, f ⊔ g) = f ⊔ complete_lattice.Inf s := le_antisymm begin @@ -676,7 +676,7 @@ le_antisymm end (le_infi $ assume g, le_infi $ assume h, sup_le_sup (le_refl f) $ Inf_le h) -lemma infi_sup_eq { f : filter α } {g : ι → filter α} : +theorem infi_sup_eq { f : filter α } {g : ι → filter α} : (⨅ x, f ⊔ g x) = f ⊔ infi g := calc (⨅ x, f ⊔ g x) = (⨅ x (h : ∃i, g i = x), f ⊔ x) : by simp; rw [infi_comm]; simp ... = f ⊔ Inf {x | ∃i, g i = x} : binfi_sup_eq @@ -685,84 +685,84 @@ calc (⨅ x, f ⊔ g x) = (⨅ x (h : ∃i, g i = x), f ⊔ x) : by simp; rw [in /- principal equations -/ @[simp] -lemma inf_principal {s t : set α} : principal s ⊓ principal t = principal (s ∩ t) := +theorem inf_principal {s t : set α} : principal s ⊓ principal t = principal (s ∩ t) := le_antisymm (by simp; exact ⟨s, subset.refl s, t, subset.refl t, subset.refl _⟩) (by simp [le_inf_iff, inter_subset_left, inter_subset_right]) @[simp] -lemma sup_principal {s t : set α} : principal s ⊔ principal t = principal (s ∪ t) := +theorem sup_principal {s t : set α} : principal s ⊔ principal t = principal (s ∪ t) := filter_eq $ set.ext $ by simp [union_subset_iff] @[simp] -lemma supr_principal {ι : Sort w} {s : ι → set α} : (⨆x, principal (s x)) = principal (⋃i, s i) := +theorem supr_principal {ι : Sort w} {s : ι → set α} : (⨆x, principal (s x)) = principal (⋃i, s i) := filter_eq $ set.ext $ assume x, by simp [supr_sets_eq]; exact (@supr_le_iff (set α) _ _ _ _).symm -lemma principal_univ : principal (univ : set α) = ⊤ := +theorem principal_univ : principal (univ : set α) = ⊤ := rfl -lemma principal_empty : principal (∅ : set α) = ⊥ := +theorem principal_empty : principal (∅ : set α) = ⊥ := rfl @[simp] -lemma principal_eq_bot_iff {s : set α} : principal s = ⊥ ↔ s = ∅ := +theorem principal_eq_bot_iff {s : set α} : principal s = ⊥ ↔ s = ∅ := ⟨assume h, principal_eq_iff_eq.mp $ by simp [principal_empty, h], assume h, by simp [*, principal_empty]⟩ @[simp] -lemma mem_pure {a : α} {s : set α} : a ∈ s → s ∈ (pure a : filter α).sets := +theorem mem_pure {a : α} {s : set α} : a ∈ s → s ∈ (pure a : filter α).sets := by simp; exact id /- map equations -/ @[simp] -lemma mem_map {f : filter α} {s : set β} {m : α → β} : +theorem mem_map {f : filter α} {s : set β} {m : α → β} : (s ∈ (map m f).sets) = ({x | m x ∈ s} ∈ f.sets) := rfl -lemma image_mem_map {f : filter α} {m : α → β} {s : set α} (hs : s ∈ f.sets): +theorem image_mem_map {f : filter α} {m : α → β} {s : set α} (hs : s ∈ f.sets): m '' s ∈ (map m f).sets := f.upwards_sets hs $ assume x hx, ⟨x, hx, rfl⟩ @[simp] -lemma map_id {f : filter α} : filter.map id f = f := +theorem map_id {f : filter α} : filter.map id f = f := filter_eq $ rfl @[simp] -lemma map_compose {γ : Type w} {f : α → β} {g : β → γ} : +theorem map_compose {γ : Type w} {f : α → β} {g : β → γ} : filter.map g ∘ filter.map f = filter.map (g ∘ f) := funext $ assume _, filter_eq $ rfl @[simp] -lemma map_sup {f g : filter α} {m : α → β} : map m (f ⊔ g) = map m f ⊔ map m g := +theorem map_sup {f g : filter α} {m : α → β} : map m (f ⊔ g) = map m f ⊔ map m g := filter_eq $ set.ext $ assume x, by simp @[simp] -lemma supr_map {ι : Sort w} {f : ι → filter α} {m : α → β} : (⨆x, map m (f x)) = map m (⨆x, f x) := +theorem supr_map {ι : Sort w} {f : ι → filter α} {m : α → β} : (⨆x, map m (f x)) = map m (⨆x, f x) := filter_eq $ set.ext $ assume x, by simp [supr_sets_eq, map] @[simp] -lemma map_bot {m : α → β} : map m ⊥ = ⊥ := +theorem map_bot {m : α → β} : map m ⊥ = ⊥ := filter_eq $ set.ext $ assume x, by simp @[simp] -lemma map_eq_bot_iff {f : filter α} {m : α → β} : map m f = ⊥ ↔ f = ⊥ := +theorem map_eq_bot_iff {f : filter α} {m : α → β} : map m f = ⊥ ↔ f = ⊥ := ⟨by rw [←empty_in_sets_eq_bot, ←empty_in_sets_eq_bot]; exact id, assume h, by simp [*]⟩ -lemma map_mono {f g : filter α} {m : α → β} (h : f ≤ g) : map m f ≤ map m g := +theorem map_mono {f g : filter α} {m : α → β} (h : f ≤ g) : map m f ≤ map m g := le_of_sup_eq $ calc map m f ⊔ map m g = map m (f ⊔ g) : map_sup ... = map m g : congr_arg (map m) $ sup_of_le_right h -lemma monotone_map {m : α → β} : monotone (map m : filter α → filter β) := +theorem monotone_map {m : α → β} : monotone (map m : filter α → filter β) := assume a b h, map_mono h -- this is a generic rule for monotone functions: -lemma map_infi_le {f : ι → filter α} {m : α → β} : +theorem map_infi_le {f : ι → filter α} {m : α → β} : map m (infi f) ≤ (⨅ i, map m (f i)) := le_infi $ assume i, map_mono $ infi_le _ _ -lemma map_infi_eq {f : ι → filter α} {m : α → β} (hf : directed (≤) f) (hι : nonempty ι) : +theorem map_infi_eq {f : ι → filter α} {m : α → β} (hf : directed (≤) f) (hι : nonempty ι) : map m (infi f) = (⨅ i, map m (f i)) := le_antisymm map_infi_le @@ -774,7 +774,7 @@ le_antisymm from infi_le_of_le i $ by simp; assumption, by simp at this; assumption) -lemma map_binfi_eq {ι : Type w} {f : ι → filter α} {m : α → β} {s : set ι} +theorem map_binfi_eq {ι : Type w} {f : ι → filter α} {m : α → β} {s : set ι} (h : directed_on (λx y, f x ≤ f y) s) (ne : ∃i, i ∈ s) : map m (⨅i∈s, f i) = (⨅i∈s, map m (f i)) := let ⟨i, hi⟩ := ne in calc map m (⨅i∈s, f i) = map m (⨅i:{i // i ∈ s}, f i.val) : by simp [infi_subtype] @@ -785,43 +785,43 @@ calc map m (⨅i∈s, f i) = map m (⨅i:{i // i ∈ s}, f i.val) : by simp [inf /- bind equations -/ -lemma mem_bind_sets {β : Type u} {s : set β} {f : filter α} {m : α → filter β} : +theorem mem_bind_sets {β : Type u} {s : set β} {f : filter α} {m : α → filter β} : s ∈ (f >>= m).sets ↔ (∃t ∈ f.sets, ∀x ∈ t, s ∈ (m x).sets) := calc s ∈ (f >>= m).sets ↔ {a | s ∈ (m a).sets} ∈ f.sets : by simp ... ↔ (∃t ∈ f.sets, t ⊆ {a | s ∈ (m a).sets}) : exists_sets_subset_iff.symm ... ↔ (∃t ∈ f.sets, ∀x ∈ t, s ∈ (m x).sets) : iff.refl _ -lemma bind_mono {β : Type u} {f : filter α} {g h : α → filter β} (h₁ : {a | g a ≤ h a} ∈ f.sets) : +theorem bind_mono {β : Type u} {f : filter α} {g h : α → filter β} (h₁ : {a | g a ≤ h a} ∈ f.sets) : f >>= g ≤ f >>= h := assume x h₂, f.upwards_sets (inter_mem_sets h₁ h₂) $ assume s ⟨gh', h'⟩, gh' h' -lemma bind_sup {β : Type u} {f g : filter α} {h : α → filter β} : +theorem bind_sup {β : Type u} {f g : filter α} {h : α → filter β} : (f ⊔ g) >>= h = (f >>= h) ⊔ (g >>= h) := by simp -lemma bind_mono2 {β : Type u} {f g : filter α} {h : α → filter β} (h₁ : f ≤ g) : +theorem bind_mono2 {β : Type u} {f g : filter α} {h : α → filter β} (h₁ : f ≤ g) : f >>= h ≤ g >>= h := assume s h', h₁ h' -lemma principal_bind {β : Type u} {s : set α} {f : α → filter β} : +theorem principal_bind {β : Type u} {s : set α} {f : α → filter β} : (principal s >>= f) = (⨆x ∈ s, f x) := show join (map f (principal s)) = (⨆x ∈ s, f x), by simp [Sup_image] -lemma seq_mono {β : Type u} {f₁ f₂ : filter (α → β)} {g₁ g₂ : filter α} +theorem seq_mono {β : Type u} {f₁ f₂ : filter (α → β)} {g₁ g₂ : filter α} (hf : f₁ ≤ f₂) (hg : g₁ ≤ g₂) : f₁ <*> g₁ ≤ f₂ <*> g₂ := le_trans (bind_mono2 hf) (bind_mono $ univ_mem_sets' $ assume f, map_mono hg) @[simp] -lemma fmap_principal {β : Type u} {s : set α} {f : α → β} : +theorem fmap_principal {β : Type u} {s : set α} {f : α → β} : f <$> principal s = principal (set.image f s) := filter_eq $ set.ext $ assume a, image_subset_iff_subset_vimage.symm -lemma mem_return_sets {a : α} {s : set α} : s ∈ (return a : filter α).sets ↔ a ∈ s := +theorem mem_return_sets {a : α} {s : set α} : s ∈ (return a : filter α).sets ↔ a ∈ s := show s ∈ (principal {a}).sets ↔ a ∈ s, by simp -lemma infi_neq_bot_of_directed {f : ι → filter α} +theorem infi_neq_bot_of_directed {f : ι → filter α} (hn : nonempty α) (hd : directed (≤) f) (hb : ∀i, f i ≠ ⊥): (infi f) ≠ ⊥ := let ⟨x⟩ := hn in assume h, have he: ∅ ∈ (infi f).sets, from h.symm ▸ mem_bot_sets, @@ -840,44 +840,44 @@ classical.by_cases end, this $ mem_univ x) -lemma infi_neq_bot_iff_of_directed {f : ι → filter α} +theorem infi_neq_bot_iff_of_directed {f : ι → filter α} (hn : nonempty α) (hd : directed (≤) f) : (infi f) ≠ ⊥ ↔ (∀i, f i ≠ ⊥) := ⟨assume neq_bot i eq_bot, neq_bot $ bot_unique $ infi_le_of_le i $ eq_bot ▸ le_refl _, infi_neq_bot_of_directed hn hd⟩ -@[simp] lemma return_neq_bot {α : Type u} {a : α} : return a ≠ (⊥ : filter α) := +@[simp] theorem return_neq_bot {α : Type u} {a : α} : return a ≠ (⊥ : filter α) := by simp [return] section vmap variables {f f₁ f₂ : filter β} {m : α → β} -lemma mem_vmap_of_mem {s : set β} (h : s ∈ f.sets) : vimage m s ∈ (vmap m f).sets := +theorem mem_vmap_of_mem {s : set β} (h : s ∈ f.sets) : vimage m s ∈ (vmap m f).sets := ⟨s, h, subset.refl _⟩ -lemma vmap_mono (h : f₁ ≤ f₂) : vmap m f₁ ≤ vmap m f₂ := +theorem vmap_mono (h : f₁ ≤ f₂) : vmap m f₁ ≤ vmap m f₂ := assume s ⟨t, ht, h_sub⟩, ⟨t, h ht, h_sub⟩ -lemma monotone_vmap : monotone (vmap m : filter β → filter α) := +theorem monotone_vmap : monotone (vmap m : filter β → filter α) := assume a b h, vmap_mono h @[simp] -lemma vmap_principal {t : set β} : vmap m (principal t) = principal (vimage m t) := +theorem vmap_principal {t : set β} : vmap m (principal t) = principal (vimage m t) := filter_eq $ set.ext $ assume s, ⟨assume ⟨u, (hu : t ⊆ u), (b : vimage m u ⊆ s)⟩, subset.trans (vimage_mono hu) b, assume : vimage m t ⊆ s, ⟨t, subset.refl t, this⟩⟩ -lemma vimage_mem_vmap {f : filter β} {m : α → β} {s : set β} (hs : s ∈ f.sets): +theorem vimage_mem_vmap {f : filter β} {m : α → β} {s : set β} (hs : s ∈ f.sets): vimage m s ∈ (vmap m f).sets := ⟨s, hs, subset.refl _⟩ -lemma le_map_vmap {f : filter β} {m : α → β} (hm : ∀x, ∃y, m y = x) : +theorem le_map_vmap {f : filter β} {m : α → β} (hm : ∀x, ∃y, m y = x) : f ≤ map m (vmap m f) := assume s ⟨t, ht, (sub : ∀x, m x ∈ t → m x ∈ s)⟩, f.upwards_sets ht $ assume x, let ⟨y, (hy : m y = x)⟩ := hm x in hy ▸ sub y -lemma vmap_map {f : filter α} {m : α → β} (h : ∀ x y, m x = m y → x = y) : +theorem vmap_map {f : filter α} {m : α → β} (h : ∀ x y, m x = m y → x = y) : vmap m (map m f) = f := have ∀s, vimage m (image m s) = s, from assume s, vimage_image_eq h, @@ -889,13 +889,13 @@ le_antisymm (assume s ⟨t, (h₁ : vimage m t ∈ f.sets), (h₂ : vimage m t ⊆ s)⟩, f.upwards_sets h₁ h₂) -lemma vmap_neq_bot {f : filter β} {m : α → β} +theorem vmap_neq_bot {f : filter β} {m : α → β} (hm : ∀t∈f.sets, ∃a, m a ∈ t) : vmap m f ≠ ⊥ := forall_sets_neq_empty_iff_neq_bot.mp $ assume s ⟨t, ht, t_s⟩, let ⟨a, (ha : a ∈ vimage m t)⟩ := hm t ht in neq_bot_of_le_neq_bot (ne_empty_of_mem ha) t_s -lemma vmap_neq_bot_of_surj {f : filter β} {m : α → β} +theorem vmap_neq_bot_of_surj {f : filter β} {m : α → β} (hf : f ≠ ⊥) (hm : ∀b, ∃a, m a = b) : vmap m f ≠ ⊥ := vmap_neq_bot $ assume t ht, let @@ -903,20 +903,20 @@ vmap_neq_bot $ assume t ht, ⟨a, (ha : m a = b)⟩ := hm b in ⟨a, ha.symm ▸ hx⟩ -lemma map_vmap_le {f : filter β} {m : α → β} : map m (vmap m f) ≤ f := +theorem map_vmap_le {f : filter β} {m : α → β} : map m (vmap m f) ≤ f := assume s hs, ⟨s, hs, subset.refl _⟩ -lemma le_vmap_map {f : filter α} {m : α → β} : f ≤ vmap m (map m f) := +theorem le_vmap_map {f : filter α} {m : α → β} : f ≤ vmap m (map m f) := assume s ⟨t, ht, h_eq⟩, f.upwards_sets ht h_eq -lemma vmap_vmap_comp {f : filter α} {m : γ → β} {n : β → α} : +theorem vmap_vmap_comp {f : filter α} {m : γ → β} {n : β → α} : vmap m (vmap n f) = vmap (n ∘ m) f := le_antisymm (assume c ⟨b, hb, (h : vimage (n ∘ m) b ⊆ c)⟩, ⟨vimage n b, vimage_mem_vmap hb, h⟩) (assume c ⟨b, ⟨a, ha, (h₁ : vimage n a ⊆ b)⟩, (h₂ : vimage m b ⊆ c)⟩, ⟨a, ha, show vimage m (vimage n a) ⊆ c, from subset.trans (vimage_mono h₁) h₂⟩) -lemma le_vmap_iff_map_le {f : filter α} {g : filter β} {m : α → β} : +theorem le_vmap_iff_map_le {f : filter α} {g : filter β} {m : α → β} : f ≤ vmap m g ↔ map m f ≤ g := ⟨assume h, le_trans (map_mono h) map_vmap_le, assume h, le_trans le_vmap_map (vmap_mono h)⟩ @@ -931,35 +931,35 @@ protected def lift (f : filter α) (g : set α → filter β) := section variables {f f₁ f₂ : filter α} {g g₁ g₂ : set α → filter β} -lemma lift_sets_eq (hg : monotone g) : (f.lift g).sets = (⋃t∈f.sets, (g t).sets) := +theorem lift_sets_eq (hg : monotone g) : (f.lift g).sets = (⋃t∈f.sets, (g t).sets) := infi_sets_eq' (assume s hs t ht, ⟨s ∩ t, inter_mem_sets hs ht, hg $ inter_subset_left s t, hg $ inter_subset_right s t⟩) ⟨univ, univ_mem_sets⟩ -lemma mem_lift {s : set β} {t : set α} (ht : t ∈ f.sets) (hs : s ∈ (g t).sets) : +theorem mem_lift {s : set β} {t : set α} (ht : t ∈ f.sets) (hs : s ∈ (g t).sets) : s ∈ (f.lift g).sets := le_principal_iff.mp $ show f.lift g ≤ principal s, from infi_le_of_le t $ infi_le_of_le ht $ le_principal_iff.mpr hs -lemma mem_lift_iff (hg : monotone g) {s : set β} : +theorem mem_lift_iff (hg : monotone g) {s : set β} : s ∈ (f.lift g).sets ↔ (∃t∈f.sets, s ∈ (g t).sets) := by rw [lift_sets_eq hg]; simp -lemma lift_mono (hf : f₁ ≤ f₂) (hg : g₁ ≤ g₂) : f₁.lift g₁ ≤ f₂.lift g₂ := +theorem lift_mono (hf : f₁ ≤ f₂) (hg : g₁ ≤ g₂) : f₁.lift g₁ ≤ f₂.lift g₂ := infi_le_infi $ assume s, infi_le_infi2 $ assume hs, ⟨hf hs, hg s⟩ -lemma lift_mono' (hg : ∀s∈f.sets, g₁ s ≤ g₂ s) : f.lift g₁ ≤ f.lift g₂ := +theorem lift_mono' (hg : ∀s∈f.sets, g₁ s ≤ g₂ s) : f.lift g₁ ≤ f.lift g₂ := infi_le_infi $ assume s, infi_le_infi $ assume hs, hg s hs -lemma map_lift_eq {m : β → γ} (hg : monotone g) : +theorem map_lift_eq {m : β → γ} (hg : monotone g) : map m (f.lift g) = f.lift (map m ∘ g) := have monotone (map m ∘ g), from monotone_comp hg monotone_map, filter_eq $ set.ext $ by simp [mem_lift_iff, hg, @mem_lift_iff _ _ f _ this] -lemma vmap_lift_eq {m : γ → β} (hg : monotone g) : +theorem vmap_lift_eq {m : γ → β} (hg : monotone g) : vmap m (f.lift g) = f.lift (vmap m ∘ g) := have monotone (vmap m ∘ g), from monotone_comp hg monotone_vmap, @@ -971,7 +971,7 @@ begin assume ⟨t₂, ht, t₁, hs, ht₁⟩, ⟨t₁, hs, t₂, ht, ht₁⟩⟩ end -lemma vmap_lift_eq2 {m : β → α} {g : set β → filter γ} (hg : monotone g) : +theorem vmap_lift_eq2 {m : β → α} {g : set β → filter γ} (hg : monotone g) : (vmap m f).lift g = f.lift (g ∘ vimage m) := le_antisymm (le_infi $ assume s, le_infi $ assume hs, @@ -979,7 +979,7 @@ le_antisymm (le_infi $ assume s, le_infi $ assume ⟨s', hs', (h_sub : vimage m s' ⊆ s)⟩, infi_le_of_le s' $ infi_le_of_le hs' $ hg h_sub) -lemma map_lift_eq2 {g : set β → filter γ} {m : α → β} (hg : monotone g) : +theorem map_lift_eq2 {g : set β → filter γ} {m : α → β} (hg : monotone g) : (map m f).lift g = f.lift (g ∘ image m) := le_antisymm (infi_le_infi2 $ assume s, ⟨image m s, @@ -992,7 +992,7 @@ le_antisymm let ⟨y, hy, h_eq⟩ := h in show x ∈ t, from h_eq ▸ hy⟩⟩) -lemma lift_comm {g : filter β} {h : set α → set β → filter γ} : +theorem lift_comm {g : filter β} {h : set α → set β → filter γ} : f.lift (λs, g.lift (h s)) = g.lift (λt, f.lift (λs, h s t)) := le_antisymm (le_infi $ assume i, le_infi $ assume hi, le_infi $ assume j, le_infi $ assume hj, @@ -1000,7 +1000,7 @@ le_antisymm (le_infi $ assume i, le_infi $ assume hi, le_infi $ assume j, le_infi $ assume hj, infi_le_of_le j $ infi_le_of_le hj $ infi_le_of_le i $ infi_le _ hi) -lemma lift_assoc {h : set β → filter γ} (hg : monotone g) : +theorem lift_assoc {h : set β → filter γ} (hg : monotone g) : (f.lift g).lift h = f.lift (λs, (g s).lift h) := le_antisymm (le_infi $ assume s, le_infi $ assume hs, le_infi $ assume t, le_infi $ assume ht, @@ -1009,11 +1009,11 @@ le_antisymm let ⟨s, hs, h'⟩ := (mem_lift_iff hg).mp ht in infi_le_of_le s $ infi_le_of_le hs $ infi_le_of_le t $ infi_le _ h') -lemma lift_lift_same_le_lift {g : set α → set α → filter β} : +theorem lift_lift_same_le_lift {g : set α → set α → filter β} : f.lift (λs, f.lift (g s)) ≤ f.lift (λs, g s s) := le_infi $ assume s, le_infi $ assume hs, infi_le_of_le s $ infi_le_of_le hs $ infi_le_of_le s $ infi_le _ hs -lemma lift_lift_same_eq_lift {g : set α → set α → filter β} +theorem lift_lift_same_eq_lift {g : set α → set α → filter β} (hg₁ : ∀s, monotone (λt, g s t)) (hg₂ : ∀t, monotone (λs, g s t)): f.lift (λs, f.lift (g s)) = f.lift (λs, g s s) := le_antisymm @@ -1024,17 +1024,17 @@ le_antisymm calc g (s ∩ t) (s ∩ t) ≤ g s (s ∩ t) : hg₂ (s ∩ t) (inter_subset_left _ _) ... ≤ g s t : hg₁ s (inter_subset_right _ _)) -lemma lift_principal {s : set α} (hg : monotone g) : +theorem lift_principal {s : set α} (hg : monotone g) : (principal s).lift g = g s := le_antisymm (infi_le_of_le s $ infi_le _ $ subset.refl _) (le_infi $ assume t, le_infi $ assume hi, hg hi) -lemma monotone_lift [weak_order γ] {f : γ → filter α} {g : γ → set α → filter β} +theorem monotone_lift [weak_order γ] {f : γ → filter α} {g : γ → set α → filter β} (hf : monotone f) (hg : monotone g) : monotone (λc, (f c).lift (g c)) := assume a b h, lift_mono (hf h) (hg h) -lemma lift_neq_bot_iff (hm : monotone g) : (f.lift g ≠ ⊥) ↔ (∀s∈f.sets, g s ≠ ⊥) := +theorem lift_neq_bot_iff (hm : monotone g) : (f.lift g ≠ ⊥) ↔ (∀s∈f.sets, g s ≠ ⊥) := classical.by_cases (assume hn : nonempty β, calc f.lift g ≠ ⊥ ↔ (⨅s : { s // s ∈ f.sets}, g s.val) ≠ ⊥ : by simp [filter.lift, infi_subtype] @@ -1058,83 +1058,83 @@ f.lift (principal ∘ h) variables {f f₁ f₂ : filter α} {h h₁ h₂ : set α → set β} -lemma mem_lift' {t : set α} (ht : t ∈ f.sets) : h t ∈ (f.lift' h).sets := +theorem mem_lift' {t : set α} (ht : t ∈ f.sets) : h t ∈ (f.lift' h).sets := le_principal_iff.mp $ show f.lift' h ≤ principal (h t), from infi_le_of_le t $ infi_le_of_le ht $ le_refl _ -lemma mem_lift'_iff (hh : monotone h) {s : set β} : s ∈ (f.lift' h).sets ↔ (∃t∈f.sets, h t ⊆ s) := +theorem mem_lift'_iff (hh : monotone h) {s : set β} : s ∈ (f.lift' h).sets ↔ (∃t∈f.sets, h t ⊆ s) := have monotone (principal ∘ h), from assume a b h, principal_mono.mpr $ hh h, by simp [filter.lift', @mem_lift_iff α β f _ this] -lemma lift'_mono (hf : f₁ ≤ f₂) (hh : h₁ ≤ h₂) : f₁.lift' h₁ ≤ f₂.lift' h₂ := +theorem lift'_mono (hf : f₁ ≤ f₂) (hh : h₁ ≤ h₂) : f₁.lift' h₁ ≤ f₂.lift' h₂ := lift_mono hf $ assume s, principal_mono.mpr $ hh s -lemma lift'_mono' (hh : ∀s∈f.sets, h₁ s ⊆ h₂ s) : f.lift' h₁ ≤ f.lift' h₂ := +theorem lift'_mono' (hh : ∀s∈f.sets, h₁ s ⊆ h₂ s) : f.lift' h₁ ≤ f.lift' h₂ := infi_le_infi $ assume s, infi_le_infi $ assume hs, principal_mono.mpr $ hh s hs -lemma lift'_cong (hh : ∀s∈f.sets, h₁ s = h₂ s) : f.lift' h₁ = f.lift' h₂ := +theorem lift'_cong (hh : ∀s∈f.sets, h₁ s = h₂ s) : f.lift' h₁ = f.lift' h₂ := le_antisymm (lift'_mono' $ assume s hs, le_of_eq $ hh s hs) (lift'_mono' $ assume s hs, le_of_eq $ (hh s hs).symm) -lemma map_lift'_eq {m : β → γ} (hh : monotone h) : map m (f.lift' h) = f.lift' (image m ∘ h) := +theorem map_lift'_eq {m : β → γ} (hh : monotone h) : map m (f.lift' h) = f.lift' (image m ∘ h) := calc map m (f.lift' h) = f.lift (map m ∘ principal ∘ h) : map_lift_eq $ monotone_comp hh monotone_principal ... = f.lift' (image m ∘ h) : by simp [function.comp, filter.lift'] -lemma map_lift'_eq2 {g : set β → set γ} {m : α → β} (hg : monotone g) : +theorem map_lift'_eq2 {g : set β → set γ} {m : α → β} (hg : monotone g) : (map m f).lift' g = f.lift' (g ∘ image m) := map_lift_eq2 $ monotone_comp hg monotone_principal -lemma vmap_lift'_eq {m : γ → β} (hh : monotone h) : +theorem vmap_lift'_eq {m : γ → β} (hh : monotone h) : vmap m (f.lift' h) = f.lift' (vimage m ∘ h) := calc vmap m (f.lift' h) = f.lift (vmap m ∘ principal ∘ h) : vmap_lift_eq $ monotone_comp hh monotone_principal ... = f.lift' (vimage m ∘ h) : by simp [function.comp, filter.lift'] -lemma vmap_lift'_eq2 {m : β → α} {g : set β → set γ} (hg : monotone g) : +theorem vmap_lift'_eq2 {m : β → α} {g : set β → set γ} (hg : monotone g) : (vmap m f).lift' g = f.lift' (g ∘ vimage m) := vmap_lift_eq2 $ monotone_comp hg monotone_principal -lemma lift'_principal {s : set α} (hh : monotone h) : +theorem lift'_principal {s : set α} (hh : monotone h) : (principal s).lift' h = principal (h s) := lift_principal $ monotone_comp hh monotone_principal -lemma principal_le_lift' {t : set β} (hh : ∀s∈f.sets, t ⊆ h s) : +theorem principal_le_lift' {t : set β} (hh : ∀s∈f.sets, t ⊆ h s) : principal t ≤ f.lift' h := le_infi $ assume s, le_infi $ assume hs, principal_mono.mpr (hh s hs) -lemma monotone_lift' [weak_order γ] {f : γ → filter α} {g : γ → set α → set β} +theorem monotone_lift' [weak_order γ] {f : γ → filter α} {g : γ → set α → set β} (hf : monotone f) (hg : monotone g) : monotone (λc, (f c).lift' (g c)) := assume a b h, lift'_mono (hf h) (hg h) -lemma lift_lift'_assoc {g : set α → set β} {h : set β → filter γ} +theorem lift_lift'_assoc {g : set α → set β} {h : set β → filter γ} (hg : monotone g) (hh : monotone h) : (f.lift' g).lift h = f.lift (λs, h (g s)) := calc (f.lift' g).lift h = f.lift (λs, (principal (g s)).lift h) : lift_assoc (monotone_comp hg monotone_principal) ... = f.lift (λs, h (g s)) : by simp [lift_principal, hh] -lemma lift'_lift'_assoc {g : set α → set β} {h : set β → set γ} +theorem lift'_lift'_assoc {g : set α → set β} {h : set β → set γ} (hg : monotone g) (hh : monotone h) : (f.lift' g).lift' h = f.lift' (λs, h (g s)) := lift_lift'_assoc hg (monotone_comp hh monotone_principal) -lemma lift'_lift_assoc {g : set α → filter β} {h : set β → set γ} +theorem lift'_lift_assoc {g : set α → filter β} {h : set β → set γ} (hg : monotone g) : (f.lift g).lift' h = f.lift (λs, (g s).lift' h) := lift_assoc hg -lemma lift_lift'_same_le_lift' {g : set α → set α → set β} : +theorem lift_lift'_same_le_lift' {g : set α → set α → set β} : f.lift (λs, f.lift' (g s)) ≤ f.lift' (λs, g s s) := lift_lift_same_le_lift -lemma lift_lift'_same_eq_lift' {g : set α → set α → set β} +theorem lift_lift'_same_eq_lift' {g : set α → set α → set β} (hg₁ : ∀s, monotone (λt, g s t)) (hg₂ : ∀t, monotone (λs, g s t)): f.lift (λs, f.lift' (g s)) = f.lift' (λs, g s s) := lift_lift_same_eq_lift (assume s, monotone_comp monotone_id $ monotone_comp (hg₁ s) monotone_principal) (assume t, monotone_comp (hg₂ t) monotone_principal) -lemma lift'_inf_principal_eq {h : set α → set β} {s : set β} : +theorem lift'_inf_principal_eq {h : set α → set β} {s : set β} : f.lift' h ⊓ principal s = f.lift' (λt, h t ∩ s) := le_antisymm (le_infi $ assume t, le_infi $ assume ht, @@ -1146,18 +1146,18 @@ le_antisymm infi_le_of_le t $ infi_le_of_le ht $ by simp; exact inter_subset_right _ _) (infi_le_of_le univ $ infi_le_of_le univ_mem_sets $ by simp; exact inter_subset_left _ _)) -lemma lift'_neq_bot_iff (hh : monotone h) : (f.lift' h ≠ ⊥) ↔ (∀s∈f.sets, h s ≠ ∅) := +theorem lift'_neq_bot_iff (hh : monotone h) : (f.lift' h ≠ ⊥) ↔ (∀s∈f.sets, h s ≠ ∅) := calc (f.lift' h ≠ ⊥) ↔ (∀s∈f.sets, principal (h s) ≠ ⊥) : lift_neq_bot_iff (monotone_comp hh monotone_principal) ... ↔ (∀s∈f.sets, h s ≠ ∅) : by simp [principal_eq_bot_iff] @[simp] -lemma lift'_id {f : filter α} : f.lift' id = f := +theorem lift'_id {f : filter α} : f.lift' id = f := le_antisymm (assume s hs, mem_lift' hs) (le_infi $ assume s, le_infi $ assume hs, by simp [hs]) -lemma le_lift' {f : filter α} {h : set α → set β} {g : filter β} +theorem le_lift' {f : filter α} {h : set α → set β} {g : filter β} (h_le : ∀s∈f.sets, h s ∈ g.sets) : g ≤ f.lift' h := le_infi $ assume s, le_infi $ assume hs, by simp [h_le]; exact h_le s hs @@ -1165,7 +1165,7 @@ end end lift -lemma vmap_eq_lift' {f : filter β} {m : α → β} : +theorem vmap_eq_lift' {f : filter β} {m : α → β} : vmap m f = f.lift' (vimage m) := filter_eq $ set.ext $ by simp [mem_lift'_iff, monotone_vimage, vmap] @@ -1191,17 +1191,17 @@ section prod protected def prod (f : filter α) (g : filter β) : filter (α × β) := f.lift $ λs, g.lift' $ λt, set.prod s t -lemma prod_mem_prod {s : set α} {t : set β} {f : filter α} {g : filter β} +theorem prod_mem_prod {s : set α} {t : set β} {f : filter α} {g : filter β} (hs : s ∈ f.sets) (ht : t ∈ g.sets) : set.prod s t ∈ (filter.prod f g).sets := le_principal_iff.mp $ show filter.prod f g ≤ principal (set.prod s t), from infi_le_of_le s $ infi_le_of_le hs $ infi_le_of_le t $ infi_le _ ht -lemma prod_same_eq {f : filter α} : filter.prod f f = f.lift' (λt, set.prod t t) := +theorem prod_same_eq {f : filter α} : filter.prod f f = f.lift' (λt, set.prod t t) := lift_lift'_same_eq_lift' (assume s, set.monotone_prod monotone_const monotone_id) (assume t, set.monotone_prod monotone_id monotone_const) -lemma mem_prod_iff {s : set (α×β)} {f : filter α} {g : filter β} : +theorem mem_prod_iff {s : set (α×β)} {f : filter α} {g : filter β} : s ∈ (filter.prod f g).sets ↔ (∃t₁∈f.sets, ∃t₂∈g.sets, set.prod t₁ t₂ ⊆ s) := begin delta filter.prod, @@ -1213,15 +1213,15 @@ begin exact (monotone_lift' monotone_const $ monotone_lam $ assume b, set.monotone_prod monotone_id monotone_const) end -lemma mem_prod_same_iff {s : set (α×α)} {f : filter α} : +theorem mem_prod_same_iff {s : set (α×α)} {f : filter α} : s ∈ (filter.prod f f).sets ↔ (∃t∈f.sets, set.prod t t ⊆ s) := by rw [prod_same_eq, mem_lift'_iff]; exact set.monotone_prod monotone_id monotone_id -lemma prod_mono {f₁ f₂ : filter α} {g₁ g₂ : filter β} (hf : f₁ ≤ f₂) (hg : g₁ ≤ g₂) : +theorem prod_mono {f₁ f₂ : filter α} {g₁ g₂ : filter β} (hf : f₁ ≤ f₂) (hg : g₁ ≤ g₂) : filter.prod f₁ g₁ ≤ filter.prod f₂ g₂ := lift_mono hf $ assume s, lift'_mono hg $ le_refl _ -lemma prod_comm {f : filter α} {g : filter β} : +theorem prod_comm {f : filter α} {g : filter β} : filter.prod f g = map (λp:β×α, (p.2, p.1)) (filter.prod g f) := eq.symm $ calc map (λp:β×α, (p.2, p.1)) (filter.prod g f) = (g.lift $ λt, map (λp:β×α, (p.2, p.1)) (f.lift' $ λs, set.prod t s)) : @@ -1231,7 +1231,7 @@ eq.symm $ calc map (λp:β×α, (p.2, p.1)) (filter.prod g f) = ... = (g.lift $ λt, f.lift' $ λs, set.prod s t) : by simp [set.image_swap_prod] ... = filter.prod f g : lift_comm -lemma prod_lift_lift {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} +theorem prod_lift_lift {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} {f₁ : filter α₁} {f₂ : filter α₂} {g₁ : set α₁ → filter β₁} {g₂ : set α₂ → filter β₂} (hg₁ : monotone g₁) (hg₂ : monotone g₂) : filter.prod (f₁.lift g₁) (f₂.lift g₂) = f₁.lift (λs, f₂.lift (λt, filter.prod (g₁ s) (g₂ t))) := @@ -1246,7 +1246,7 @@ begin exact hg₁ end -lemma prod_lift'_lift' {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} +theorem prod_lift'_lift' {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} {f₁ : filter α₁} {f₂ : filter α₂} {g₁ : set α₁ → set β₁} {g₂ : set α₂ → set β₂} (hg₁ : monotone g₁) (hg₂ : monotone g₂) : filter.prod (f₁.lift' g₁) (f₂.lift' g₂) = f₁.lift (λs, f₂.lift' (λt, set.prod (g₁ s) (g₂ t))) := @@ -1262,7 +1262,7 @@ begin assume x, set.monotone_prod monotone_id monotone_const) end -lemma prod_map_map_eq {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} +theorem prod_map_map_eq {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} {f₁ : filter α₁} {f₂ : filter α₂} {m₁ : α₁ → β₁} {m₂ : α₂ → β₂} : filter.prod (map m₁ f₁) (map m₂ f₂) = map (λp:α₁×α₂, (m₁ p.1, m₂ p.2)) (filter.prod f₁ f₂) := begin @@ -1279,7 +1279,7 @@ begin exact set.prod_image_image_eq end -lemma prod_vmap_vmap_eq {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} +theorem prod_vmap_vmap_eq {α₁ : Type u} {α₂ : Type v} {β₁ : Type w} {β₂ : Type x} {f₁ : filter α₁} {f₂ : filter α₂} {m₁ : β₁ → α₁} {m₂ : β₂ → α₂} : filter.prod (vmap m₁ f₁) (vmap m₂ f₂) = vmap (λp:β₁×β₂, (m₁ p.1, m₂ p.2)) (filter.prod f₁ f₂) := have ∀s t, set.prod (vimage m₁ s) (vimage m₂ t) = vimage (λp:β₁×β₂, (m₁ p.1, m₂ p.2)) (set.prod s t), @@ -1297,7 +1297,7 @@ begin exact monotone_vimage end -lemma prod_inf_prod {f₁ f₂ : filter α} {g₁ g₂ : filter β} : +theorem prod_inf_prod {f₁ f₂ : filter α} {g₁ g₂ : filter β} : filter.prod f₁ g₁ ⊓ filter.prod f₂ g₂ = filter.prod (f₁ ⊓ f₂) (g₁ ⊓ g₂) := le_antisymm (le_infi $ assume s, le_infi $ assume hs, le_infi $ assume t, le_infi $ assume ht, @@ -1310,7 +1310,7 @@ le_antisymm end) (le_inf (prod_mono inf_le_left inf_le_left) (prod_mono inf_le_right inf_le_right)) -lemma prod_neq_bot {f : filter α} {g : filter β} : +theorem prod_neq_bot {f : filter α} {g : filter β} : filter.prod f g ≠ ⊥ ↔ (f ≠ ⊥ ∧ g ≠ ⊥) := calc filter.prod f g ≠ ⊥ ↔ (∀s∈f.sets, g.lift' (set.prod s) ≠ ⊥) : begin @@ -1334,7 +1334,7 @@ calc filter.prod f g ≠ ⊥ ↔ (∀s∈f.sets, g.lift' (set.prod s) ≠ ⊥) : assume ⟨h₁, h₂⟩ s hs t ht, ⟨h₁ s hs, h₂ t ht⟩⟩ ... ↔ _ : by simp [forall_sets_neq_empty_iff_neq_bot] -lemma prod_principal_principal {s : set α} {t : set β} : +theorem prod_principal_principal {s : set α} {t : set β} : filter.prod (principal s) (principal t) = principal (set.prod s t) := begin delta filter.prod, @@ -1346,15 +1346,15 @@ end end prod -lemma mem_infi_sets {f : ι → filter α} (i : ι) : ∀{s}, s ∈ (f i).sets → s ∈ (⨅i, f i).sets := +theorem mem_infi_sets {f : ι → filter α} (i : ι) : ∀{s}, s ∈ (f i).sets → s ∈ (⨅i, f i).sets := show (⨅i, f i) ≤ f i, from infi_le _ _ @[simp] -lemma mem_top_sets_iff {s : set α} : s ∈ (⊤ : filter α).sets ↔ s = univ := +theorem mem_top_sets_iff {s : set α} : s ∈ (⊤ : filter α).sets ↔ s = univ := ⟨assume h, top_unique $ h, assume h, h.symm ▸ univ_mem_sets⟩ @[elab_as_eliminator] -lemma infi_sets_induct {f : ι → filter α} {s : set α} (hs : s ∈ (infi f).sets) {p : set α → Prop} +theorem infi_sets_induct {f : ι → filter α} {s : set α} (hs : s ∈ (infi f).sets) {p : set α → Prop} (uni : p univ) (ins : ∀{i s₁ s₂}, s₁ ∈ (f i).sets → p s₂ → p (s₁ ∩ s₂)) (upw : ∀{s₁ s₂}, s₁ ⊆ s₂ → p s₁ → p s₂) : p s := @@ -1379,7 +1379,7 @@ begin show p s, from upw hs $ ins hs₁ hs₂ } end -lemma lift_infi {f : ι → filter α} {g : set α → filter β} +theorem lift_infi {f : ι → filter α} {g : set α → filter β} (hι : nonempty ι) (hg : ∀{s t}, g s ⊓ g t = g (s ∩ t)) : (infi f).lift g = (⨅i, (f i).lift g) := le_antisymm (le_infi $ assume i, lift_mono (infi_le _ _) (le_refl _)) @@ -1394,7 +1394,7 @@ le_antisymm (assume s₁ s₂ hs₁ hs₂, le_trans hs₂ $ g_mono hs₁), by rw [lift_sets_eq g_mono]; simp; exact assume ⟨t, hs, ht⟩, this t ht hs) -lemma lift_infi' {f : ι → filter α} {g : set α → filter β} +theorem lift_infi' {f : ι → filter α} {g : set α → filter β} (hι : nonempty ι) (hf : directed (≤) f) (hg : monotone g) : (infi f).lift g = (⨅i, (f i).lift g) := le_antisymm (le_infi $ assume i, lift_mono (infi_le _ _) (le_refl _)) @@ -1405,11 +1405,11 @@ le_antisymm exact assume ⟨t, hs, i, ht⟩, mem_infi_sets i $ mem_lift ht hs end) -lemma lift'_infi {f : ι → filter α} {g : set α → set β} +theorem lift'_infi {f : ι → filter α} {g : set α → set β} (hι : nonempty ι) (hg : ∀{s t}, g s ∩ g t = g (s ∩ t)) : (infi f).lift' g = (⨅i, (f i).lift' g) := lift_infi hι $ by simp; apply assume s t, hg -lemma map_eq_vmap_of_inverse {f : filter α} {m : α → β} {n : β → α} +theorem map_eq_vmap_of_inverse {f : filter α} {m : α → β} {n : β → α} (h₁ : m ∘ n = id) (h₂ : n ∘ m = id) : map m f = vmap n f := le_antisymm (assume b ⟨a, ha, (h : vimage n a ⊆ b)⟩, f.upwards_sets ha $ @@ -1418,7 +1418,7 @@ le_antisymm (assume b (hb : vimage m b ∈ f.sets), ⟨vimage m b, hb, show vimage (m ∘ n) b ⊆ b, by simp [h₁]; apply subset.refl⟩) -lemma map_swap_vmap_swap_eq {f : filter (α × β)} : prod.swap <$> f = vmap prod.swap f := +theorem map_swap_vmap_swap_eq {f : filter (α × β)} : prod.swap <$> f = vmap prod.swap f := map_eq_vmap_of_inverse prod.swap_swap_eq prod.swap_swap_eq /- towards -/ @@ -1435,7 +1435,7 @@ variables {f g : filter α} def ultrafilter (f : filter α) := f ≠ ⊥ ∧ ∀g, g ≠ ⊥ → g ≤ f → f ≤ g -lemma ultrafilter_pure {a : α} : ultrafilter (pure a) := +theorem ultrafilter_pure {a : α} : ultrafilter (pure a) := ⟨return_neq_bot, assume g hg ha, have {a} ∈ g.sets, by simp at ha; assumption, @@ -1450,10 +1450,10 @@ lemma ultrafilter_pure {a : α} : ultrafilter (pure a) := hg this end⟩ -lemma ultrafilter_unique (hg : ultrafilter g) (hf : f ≠ ⊥) (h : f ≤ g) : f = g := +theorem ultrafilter_unique (hg : ultrafilter g) (hf : f ≠ ⊥) (h : f ≤ g) : f = g := le_antisymm h (hg.right _ hf h) -lemma exists_ultrafilter (h : f ≠ ⊥) : ∃u, u ≤ f ∧ ultrafilter u := +theorem exists_ultrafilter (h : f ≠ ⊥) : ∃u, u ≤ f ∧ ultrafilter u := let τ := {f' // f' ≠ ⊥ ∧ f' ≤ f}, r : τ → τ → Prop := λt₁ t₂, t₂.val ≤ t₁.val, @@ -1474,24 +1474,24 @@ let ⟨uτ, hmin⟩ := this in ⟨uτ.val, uτ.property.right, uτ.property.left, assume g hg₁ hg₂, hmin ⟨g, hg₁, le_trans hg₂ uτ.property.right⟩ hg₂⟩ -lemma le_of_ultrafilter {g : filter α} (hf : ultrafilter f) (h : f ⊓ g ≠ ⊥) : +theorem le_of_ultrafilter {g : filter α} (hf : ultrafilter f) (h : f ⊓ g ≠ ⊥) : f ≤ g := le_of_inf_eq $ ultrafilter_unique hf h inf_le_left -lemma mem_or_compl_mem_of_ultrafilter (hf : ultrafilter f) (s : set α) : +theorem mem_or_compl_mem_of_ultrafilter (hf : ultrafilter f) (s : set α) : s ∈ f.sets ∨ - s ∈ f.sets := or_of_not_implies' $ assume : - s ∉ f.sets, have f ≤ principal s, from le_of_ultrafilter hf $ assume h, this $ mem_sets_of_neq_bot $ by simp [*], by simp at this; assumption -lemma mem_or_mem_of_ultrafilter {s t : set α} (hf : ultrafilter f) (h : s ∪ t ∈ f.sets) : +theorem mem_or_mem_of_ultrafilter {s t : set α} (hf : ultrafilter f) (h : s ∪ t ∈ f.sets) : s ∈ f.sets ∨ t ∈ f.sets := (mem_or_compl_mem_of_ultrafilter hf s).imp_right (assume : -s ∈ f.sets, f.upwards_sets (inter_mem_sets this h) $ assume x ⟨hnx, hx⟩, hx.resolve_left hnx) -lemma mem_of_finite_sUnion_ultrafilter {s : set (set α)} (hf : ultrafilter f) (hs : finite s) +theorem mem_of_finite_sUnion_ultrafilter {s : set (set α)} (hf : ultrafilter f) (hs : finite s) : ⋃₀ s ∈ f.sets → ∃t∈s, t ∈ f.sets := begin induction hs, @@ -1503,21 +1503,21 @@ begin (assume h, let ⟨t, hts', ht⟩ := ih h in ⟨t, ht, or.inr hts'⟩) } end -lemma mem_of_finite_Union_ultrafilter {is : set β} {s : β → set α} +theorem mem_of_finite_Union_ultrafilter {is : set β} {s : β → set α} (hf : ultrafilter f) (his : finite is) (h : (⋃i∈is, s i) ∈ f.sets) : ∃i∈is, s i ∈ f.sets := have his : finite (image s is), from finite_image his, have h : (⋃₀ image s is) ∈ f.sets, from by simp [sUnion_image]; assumption, let ⟨t, ⟨i, hi, h_eq⟩, (ht : t ∈ f.sets)⟩ := mem_of_finite_sUnion_ultrafilter hf his h in ⟨i, hi, h_eq.symm ▸ ht⟩ -lemma ultrafilter_of_split {f : filter α} (hf : f ≠ ⊥) (h : ∀s, s ∈ f.sets ∨ - s ∈ f.sets) : +theorem ultrafilter_of_split {f : filter α} (hf : f ≠ ⊥) (h : ∀s, s ∈ f.sets ∨ - s ∈ f.sets) : ultrafilter f := ⟨hf, assume g hg g_le s hs, (h s).elim id $ assume : - s ∈ f.sets, have s ∩ -s ∈ g.sets, from inter_mem_sets hs (g_le this), by simp [empty_in_sets_eq_bot, hg] at this; contradiction⟩ -lemma ultrafilter_map {f : filter α} {m : α → β} (h : ultrafilter f) : ultrafilter (map m f) := +theorem ultrafilter_map {f : filter α} {m : α → β} (h : ultrafilter f) : ultrafilter (map m f) := ultrafilter_of_split (by simp [map_eq_bot_iff, h.left]) $ assume s, show vimage m s ∈ f.sets ∨ - vimage m s ∈ f.sets, from mem_or_compl_mem_of_ultrafilter h (vimage m s) @@ -1525,7 +1525,7 @@ ultrafilter_of_split (by simp [map_eq_bot_iff, h.left]) $ noncomputable def ultrafilter_of (f : filter α) : filter α := if h : f = ⊥ then ⊥ else classical.epsilon (λu, u ≤ f ∧ ultrafilter u) -lemma ultrafilter_of_spec (h : f ≠ ⊥) : ultrafilter_of f ≤ f ∧ ultrafilter (ultrafilter_of f) := +theorem ultrafilter_of_spec (h : f ≠ ⊥) : ultrafilter_of f ≤ f ∧ ultrafilter (ultrafilter_of f) := begin have h' := classical.epsilon_spec (exists_ultrafilter h), simp [ultrafilter_of, dif_neg, h], @@ -1533,14 +1533,14 @@ begin assumption end -lemma ultrafilter_of_le : ultrafilter_of f ≤ f := +theorem ultrafilter_of_le : ultrafilter_of f ≤ f := if h : f = ⊥ then by simp [ultrafilter_of, dif_pos, h]; exact le_refl _ else (ultrafilter_of_spec h).left -lemma ultrafilter_ultrafilter_of (h : f ≠ ⊥) : ultrafilter (ultrafilter_of f) := +theorem ultrafilter_ultrafilter_of (h : f ≠ ⊥) : ultrafilter (ultrafilter_of f) := (ultrafilter_of_spec h).right -lemma ultrafilter_of_ultrafilter (h : ultrafilter f) : ultrafilter_of f = f := +theorem ultrafilter_of_ultrafilter (h : ultrafilter f) : ultrafilter_of f = f := ultrafilter_unique h (ultrafilter_ultrafilter_of h.left).left ultrafilter_of_le end ultrafilter diff --git a/algebra/lattice/fixed_points.lean b/algebra/lattice/fixed_points.lean index a95c02426b4ac..0cf63703ee971 100644 --- a/algebra/lattice/fixed_points.lean +++ b/algebra/lattice/fixed_points.lean @@ -11,7 +11,7 @@ import algebra.lattice.complete_lattice universes u v w variables {α : Type u} {β : Type v} {γ : Type w} -lemma ge_of_eq [weak_order α] {a b : α} : a = b → a ≥ b := +theorem ge_of_eq [weak_order α] {a b : α} : a = b → a ≥ b := λ h, h ▸ le_refl a namespace lattice @@ -22,18 +22,18 @@ variables [complete_lattice α] {f : α → α} def lfp (f : α → α) : α := Inf {a | f a ≤ a} def gfp (f : α → α) : α := Sup {a | a ≤ f a} -lemma lfp_le {a : α} (h : f a ≤ a) : lfp f ≤ a := +theorem lfp_le {a : α} (h : f a ≤ a) : lfp f ≤ a := Inf_le h -lemma le_lfp {a : α} (h : ∀b, f b ≤ b → a ≤ b) : a ≤ lfp f := +theorem le_lfp {a : α} (h : ∀b, f b ≤ b → a ≤ b) : a ≤ lfp f := le_Inf h -lemma lfp_eq (m : monotone f) : lfp f = f (lfp f) := +theorem lfp_eq (m : monotone f) : lfp f = f (lfp f) := have f (lfp f) ≤ lfp f, from le_lfp $ assume b, assume h : f b ≤ b, le_trans (m (lfp_le h)) h, le_antisymm (lfp_le (m this)) this -lemma lfp_induct {p : α → Prop} (m : monotone f) +theorem lfp_induct {p : α → Prop} (m : monotone f) (step : ∀a, p a → a ≤ lfp f → p (f a)) (sup : ∀s, (∀a∈s, p a) → p (Sup s)) : p (lfp f) := let s := {a | a ≤ lfp f ∧ p a} in @@ -46,21 +46,21 @@ have Sup s = lfp f, ⟨le_trans (m this) $ ge_of_eq $ lfp_eq m, step _ p_s this⟩, this ▸ p_s -lemma monotone_lfp : monotone (@lfp α _) := +theorem monotone_lfp : monotone (@lfp α _) := assume f g, assume : f ≤ g, le_lfp $ assume a, assume : g a ≤ a, lfp_le $ le_trans (‹f ≤ g› a) this -lemma le_gfp {a : α} (h : a ≤ f a) : a ≤ gfp f := +theorem le_gfp {a : α} (h : a ≤ f a) : a ≤ gfp f := le_Sup h -lemma gfp_le {a : α} (h : ∀b, b ≤ f b → b ≤ a) : gfp f ≤ a := +theorem gfp_le {a : α} (h : ∀b, b ≤ f b → b ≤ a) : gfp f ≤ a := Sup_le h -lemma gfp_eq (m : monotone f) : gfp f = f (gfp f) := +theorem gfp_eq (m : monotone f) : gfp f = f (gfp f) := have gfp f ≤ f (gfp f), from gfp_le $ assume b, assume h : b ≤ f b, le_trans h (m (le_gfp h)), le_antisymm this (le_gfp (m this)) -lemma gfp_induct {p : α → Prop} (m : monotone f) +theorem gfp_induct {p : α → Prop} (m : monotone f) (step : ∀a, p a → gfp f ≤ a → p (f a)) (inf : ∀s, (∀a∈s, p a) → p (Inf s)) : p (gfp f) := let s := {a | gfp f ≤ a ∧ p a} in @@ -73,7 +73,7 @@ have Inf s = gfp f, ⟨le_trans (le_of_eq $ gfp_eq m) (m this), step _ p_s this⟩) this, this ▸ p_s -lemma monotone_gfp : monotone (@gfp α _) := +theorem monotone_gfp : monotone (@gfp α _) := assume f g, assume : f ≤ g, gfp_le $ assume a, assume : a ≤ f a, le_gfp $ le_trans this (‹f ≤ g› a) end fixedpoint @@ -82,20 +82,20 @@ section fixedpoint_eqn variables [complete_lattice α] [complete_lattice β] {f : β → α} {g : α → β} -- Rolling rule -lemma lfp_comp (m_f : monotone f) (m_g : monotone g) : lfp (f ∘ g) = f (lfp (g ∘ f)) := +theorem lfp_comp (m_f : monotone f) (m_g : monotone g) : lfp (f ∘ g) = f (lfp (g ∘ f)) := le_antisymm (lfp_le $ m_f $ ge_of_eq $ lfp_eq $ monotone_comp m_f m_g) (le_lfp $ assume a fg_le, le_trans (m_f $ lfp_le $ show (g ∘ f) (g a) ≤ g a, from m_g fg_le) fg_le) -lemma gfp_comp (m_f : monotone f) (m_g : monotone g) : gfp (f ∘ g) = f (gfp (g ∘ f)) := +theorem gfp_comp (m_f : monotone f) (m_g : monotone g) : gfp (f ∘ g) = f (gfp (g ∘ f)) := le_antisymm (gfp_le $ assume a fg_le, le_trans fg_le $ m_f $ le_gfp $ show g a ≤ (g ∘ f) (g a), from m_g fg_le) (le_gfp $ m_f $ le_of_eq $ gfp_eq $ monotone_comp m_f m_g) -- Diagonal rule -lemma lfp_lfp {h : α → α → α} (m : ∀⦃a b c d⦄, a ≤ b → c ≤ d → h a c ≤ h b d) : +theorem lfp_lfp {h : α → α → α} (m : ∀⦃a b c d⦄, a ≤ b → c ≤ d → h a c ≤ h b d) : lfp (lfp ∘ h) = lfp (λx, h x x) := let f := lfp (lfp ∘ h) in have f_eq : f = lfp (h f), @@ -107,7 +107,7 @@ le_antisymm ... = h f (lfp (h f)) : lfp_eq $ assume a b h, m (le_refl _) h ... = h f f : congr_arg (h f) f_eq.symm) -lemma gfp_gfp {h : α → α → α} (m : ∀⦃a b c d⦄, a ≤ b → c ≤ d → h a c ≤ h b d) : +theorem gfp_gfp {h : α → α → α} (m : ∀⦃a b c d⦄, a ≤ b → c ≤ d → h a c ≤ h b d) : gfp (gfp ∘ h) = gfp (λx, h x x) := let f := gfp (gfp ∘ h) in have f_eq : f = gfp (h f), diff --git a/algebra/lattice/zorn.lean b/algebra/lattice/zorn.lean index 52f787e7b59fa..2c8b135888a5a 100644 --- a/algebra/lattice/zorn.lean +++ b/algebra/lattice/zorn.lean @@ -23,7 +23,7 @@ local infix ` ≺ `:50 := r def chain (c : set α) := pairwise_on c (λx y, x ≺ y ∨ y ≺ x) -lemma chain_insert {c : set α} {a : α} (hc : chain c) (ha : ∀b∈c, b ≠ a → a ≺ b ∨ b ≺ a) : +theorem chain_insert {c : set α} {a : α} (hc : chain c) (ha : ∀b∈c, b ≠ a → a ≺ b ∨ b ≺ a) : chain (insert a c) := forall_insert_of_forall (assume x hx, forall_insert_of_forall (hc x hx) (assume hneq, (ha x hx hneq).symm)) @@ -36,20 +36,20 @@ def is_max_chain (c : set α) := chain c ∧ ¬ (∃c', super_chain c c') def succ_chain (c : set α) := if h : ∃c', chain c ∧ super_chain c c' then some h else c -lemma succ_spec {c : set α} (h : ∃c', chain c ∧ super_chain c c') : +theorem succ_spec {c : set α} (h : ∃c', chain c ∧ super_chain c c') : super_chain c (succ_chain c) := let ⟨c', hc'⟩ := h in have chain c ∧ super_chain c (some h), from @some_spec _ (λc', chain c ∧ super_chain c c') _, by simp [succ_chain, dif_pos, h, this.right] -lemma chain_succ {c : set α} (hc : chain c) : chain (succ_chain c) := +theorem chain_succ {c : set α} (hc : chain c) : chain (succ_chain c) := if h : ∃c', chain c ∧ super_chain c c' then (succ_spec h).left else by simp [succ_chain, dif_neg, h]; exact hc -lemma super_of_not_max {c : set α} (hc₁ : chain c) (hc₂ : ¬ is_max_chain c) : +theorem super_of_not_max {c : set α} (hc₁ : chain c) (hc₂ : ¬ is_max_chain c) : super_chain c (succ_chain c) := begin simp [is_max_chain, not_and_iff, not_not_iff] at hc₂, @@ -59,7 +59,7 @@ begin from succ_spec ⟨c', hc₁, hc'⟩ end -lemma succ_increasing {c : set α} : c ⊆ succ_chain c := +theorem succ_increasing {c : set α} : c ⊆ succ_chain c := if h : ∃c', chain c ∧ super_chain c c' then have super_chain c (succ_chain c), from succ_spec h, this.right.left @@ -69,12 +69,12 @@ inductive chain_closure : set α → Prop | succ : ∀{s}, chain_closure s → chain_closure (succ_chain s) | union : ∀{s}, (∀a∈s, chain_closure a) → chain_closure (⋃₀ s) -lemma chain_closure_empty : chain_closure ∅ := +theorem chain_closure_empty : chain_closure ∅ := have chain_closure (⋃₀ ∅), from chain_closure.union $ assume a h, h.rec _, by simp at this; assumption -lemma chain_closure_closure : chain_closure (⋃₀ chain_closure) := +theorem chain_closure_closure : chain_closure (⋃₀ chain_closure) := chain_closure.union $ assume s hs, hs variables {c c₁ c₂ c₃ : set α} @@ -125,12 +125,12 @@ begin { exact (h₂ $ subset.trans succ_increasing h) } } end -lemma chain_closure_total (hc₁ : chain_closure c₁) (hc₂ : chain_closure c₂) : c₁ ⊆ c₂ ∨ c₂ ⊆ c₁ := +theorem chain_closure_total (hc₁ : chain_closure c₁) (hc₂ : chain_closure c₂) : c₁ ⊆ c₂ ∨ c₂ ⊆ c₁ := have c₁ ⊆ c₂ ∨ succ_chain c₂ ⊆ c₁, from chain_closure_succ_total_aux hc₁ hc₂ $ assume c₃ hc₃, chain_closure_succ_total hc₃ hc₂, or.imp_right (assume : succ_chain c₂ ⊆ c₁, subset.trans succ_increasing this) this -lemma chain_closure_succ_fixpoint (hc₁ : chain_closure c₁) (hc₂ : chain_closure c₂) +theorem chain_closure_succ_fixpoint (hc₁ : chain_closure c₁) (hc₂ : chain_closure c₂) (h_eq : succ_chain c₂ = c₂) : c₁ ⊆ c₂ := begin induction hc₁, @@ -141,7 +141,7 @@ begin exact (sUnion_subset $ assume c₁ hc₁, ih c₁ hc₁) } end -lemma chain_closure_succ_fixpoint_iff (hc : chain_closure c) : +theorem chain_closure_succ_fixpoint_iff (hc : chain_closure c) : succ_chain c = c ↔ c = ⋃₀ chain_closure := ⟨assume h, subset.antisymm (subset_sUnion_of_mem hc) @@ -153,7 +153,7 @@ lemma chain_closure_succ_fixpoint_iff (hc : chain_closure c) : ... = c : this.symm) succ_increasing⟩ -lemma chain_chain_closure (hc : chain_closure c) : chain c := +theorem chain_chain_closure (hc : chain_closure c) : chain c := begin induction hc, case _root_.zorn.chain_closure.succ c hc h { @@ -173,7 +173,7 @@ def max_chain := ⋃₀ chain_closure There exists a maximal totally ordered subset of `α`. Note that we do not require `α` to be partially ordered by `r`. -/ -lemma max_chain_spec : is_max_chain max_chain := +theorem max_chain_spec : is_max_chain max_chain := classical.by_contradiction $ assume : ¬ is_max_chain (⋃₀ chain_closure), have super_chain (⋃₀ chain_closure) (succ_chain (⋃₀ chain_closure)), @@ -186,7 +186,7 @@ h₃ this.symm /-- Zorn's lemma If every chain has an upper bound, then there is a maximal element -/ -lemma zorn (h : ∀c, chain c → ∃ub, ∀a∈c, a ≺ ub) (trans : ∀{a b c}, a ≺ b → b ≺ c → a ≺ c) : +theorem zorn (h : ∀c, chain c → ∃ub, ∀a∈c, a ≺ ub) (trans : ∀{a b c}, a ≺ b → b ≺ c → a ≺ c) : ∃m, ∀a, m ≺ a → a ≺ m := have ∃ub, ∀a∈max_chain, a ≺ ub, from h _ $ max_chain_spec.left, @@ -201,7 +201,7 @@ let ⟨ub, (hub : ∀a∈max_chain, a ≺ ub)⟩ := this in end chain -lemma zorn_weak_order {α : Type u} [weak_order α] +theorem zorn_weak_order {α : Type u} [weak_order α] (h : ∀c:set α, @chain α (≤) c → ∃ub, ∀a∈c, a ≤ ub) : ∃m:α, ∀a, m ≤ a → a = m := let ⟨m, hm⟩ := @zorn α (≤) h (assume a b c, le_trans) in ⟨m, assume a ha, le_antisymm (hm a ha) ha⟩ diff --git a/algebra/order.lean b/algebra/order.lean index 7d1f55f757d83..5255ffb1d3135 100644 --- a/algebra/order.lean +++ b/algebra/order.lean @@ -15,11 +15,11 @@ variables [weak_order α] [weak_order β] [weak_order γ] def monotone (f : α → β) := ∀{{a b}}, a ≤ b → f a ≤ f b -lemma monotone_id : @monotone α α _ _ id := assume x y h, h +theorem monotone_id : @monotone α α _ _ id := assume x y h, h -lemma monotone_const {b : β} : monotone (λ(a:α), b) := assume x y h, le_refl b +theorem monotone_const {b : β} : monotone (λ(a:α), b) := assume x y h, le_refl b -lemma monotone_comp {f : α → β} {g : β → γ} (m_f : monotone f) (m_g : monotone g) : +theorem monotone_comp {f : α → β} {g : β → γ} (m_f : monotone f) (m_g : monotone g) : monotone (g ∘ f) := assume a b h, m_g (m_f h) @@ -61,22 +61,22 @@ def weak_order_dual (wo : weak_order α) : weak_order α := le_trans := assume a b c h₁ h₂, le_trans h₂ h₁, le_antisymm := assume a b h₁ h₂, le_antisymm h₂ h₁ } -lemma le_dual_eq_le {α : Type} (wo : weak_order α) (a b : α) : +theorem le_dual_eq_le {α : Type} (wo : weak_order α) (a b : α) : @has_le.le _ (@weak_order.to_has_le _ (weak_order_dual wo)) a b = @has_le.le _ (@weak_order.to_has_le _ wo) b a := rfl -lemma comp_le_comp_left_of_monotone [weak_order α] [weak_order β] [weak_order γ] +theorem comp_le_comp_left_of_monotone [weak_order α] [weak_order β] [weak_order γ] {f : β → α} {g h : γ → β} (m_f : monotone f) (le_gh : g ≤ h) : has_le.le.{max w u} (f ∘ g) (f ∘ h) := assume x, m_f (le_gh x) section monotone variables [weak_order α] [weak_order γ] -lemma monotone_lam {f : α → β → γ} (m : ∀b, monotone (λa, f a b)) : monotone f := +theorem monotone_lam {f : α → β → γ} (m : ∀b, monotone (λa, f a b)) : monotone f := assume a a' h b, m b h -lemma monotone_app (f : β → α → γ) (b : β) (m : monotone (λa b, f b a)) : monotone (f b) := +theorem monotone_app (f : β → α → γ) (b : β) (m : monotone (λa b, f b a)) : monotone (f b) := assume a a' h, m h b end monotone diff --git a/data/bitvec.lean b/data/bitvec.lean index 23ccae181c251..c3b37212487d3 100644 --- a/data/bitvec.lean +++ b/data/bitvec.lean @@ -155,7 +155,7 @@ section conversion protected def to_nat {n : nat} (v : bitvec n) : nat := bits_to_nat (to_list v) - lemma bits_to_nat_to_list {n : ℕ} (x : bitvec n) + theorem bits_to_nat_to_list {n : ℕ} (x : bitvec n) : bitvec.to_nat x = bits_to_nat (vector.to_list x) := rfl theorem to_nat_append {m : ℕ} (xs : bitvec m) (b : bool) @@ -180,7 +180,7 @@ section conversion simp [cond_to_bool_mod_two], end - lemma of_nat_succ {k n : ℕ} + theorem of_nat_succ {k n : ℕ} : bitvec.of_nat (succ k) n = bitvec.of_nat k (n / 2) ++ₜ[to_bool (n % 2 = 1)] := rfl diff --git a/data/bool.lean b/data/bool.lean index 401e8c076ec6f..5b102aa413b3e 100644 --- a/data/bool.lean +++ b/data/bool.lean @@ -133,44 +133,44 @@ namespace bool | tt tt := ff @[simp] - lemma ff_bxor_ff : bxor ff ff = ff := rfl + theorem ff_bxor_ff : bxor ff ff = ff := rfl @[simp] - lemma ff_bxor_tt : bxor ff tt = tt := rfl + theorem ff_bxor_tt : bxor ff tt = tt := rfl @[simp] - lemma tt_bxor_ff : bxor tt ff = tt := rfl + theorem tt_bxor_ff : bxor tt ff = tt := rfl @[simp] - lemma tt_bxor_tt : bxor tt tt = ff := rfl + theorem tt_bxor_tt : bxor tt tt = ff := rfl @[simp] - lemma bxor_self (a : bool) : bxor a a = ff := + theorem bxor_self (a : bool) : bxor a a = ff := by cases a; simp @[simp] - lemma bxor_ff (a : bool) : bxor a ff = a := + theorem bxor_ff (a : bool) : bxor a ff = a := by cases a; simp @[simp] - lemma bxor_tt (a : bool) : bxor a tt = bnot a := + theorem bxor_tt (a : bool) : bxor a tt = bnot a := by cases a; simp @[simp] - lemma ff_bxor (a : bool) : bxor ff a = a := + theorem ff_bxor (a : bool) : bxor ff a = a := by cases a; simp @[simp] - lemma tt_bxor (a : bool) : bxor tt a = bnot a := + theorem tt_bxor (a : bool) : bxor tt a = bnot a := by cases a; simp @[simp] - lemma bxor_comm (a b : bool) : bxor a b = bxor b a := + theorem bxor_comm (a b : bool) : bxor a b = bxor b a := by cases a; simp @[simp] - lemma bxor_assoc (a b c : bool) : bxor (bxor a b) c = bxor a (bxor b c) := + theorem bxor_assoc (a b c : bool) : bxor (bxor a b) c = bxor a (bxor b c) := by cases a; cases b; simp @[simp] - lemma bxor_left_comm (a b c : bool) : bxor a (bxor b c) = bxor b (bxor a c) := + theorem bxor_left_comm (a b c : bool) : bxor a (bxor b c) = bxor b (bxor a c) := by cases a; cases b; simp instance forall_decidable {P : bool → Prop} [decidable_pred P] : decidable (∀b, P b) := diff --git a/data/fin.lean b/data/fin.lean index de1e5fa75e922..7ccfc43b00800 100644 --- a/data/fin.lean +++ b/data/fin.lean @@ -1,10 +1,10 @@ open fin nat -lemma lt_succ_of_lt {a b : nat} (h : a < b) : a < b + 1 := lt_add_of_lt_of_pos h one_pos +theorem lt_succ_of_lt {a b : nat} (h : a < b) : a < b + 1 := lt_add_of_lt_of_pos h one_pos def raise_fin {n : ℕ} (k : fin n) : fin (n + 1) := ⟨val k, lt_succ_of_lt (is_lt k)⟩ -lemma eq_of_lt_succ_of_not_lt {a b : ℕ} (h1 : a < b + 1) (h2 : ¬ a < b) : a = b := +theorem eq_of_lt_succ_of_not_lt {a b : ℕ} (h1 : a < b + 1) (h2 : ¬ a < b) : a = b := have h3 : a ≤ b, from le_of_lt_succ h1, or.elim (eq_or_lt_of_not_lt h2) (λ h, h) (λ h, absurd h (not_lt_of_ge h3)) diff --git a/data/hash_map.lean b/data/hash_map.lean index 2831c4df4b8bc..d4a8196893024 100644 --- a/data/hash_map.lean +++ b/data/hash_map.lean @@ -54,7 +54,7 @@ end def foldl {δ : Type w} (d : δ) (f : δ → Π a, β a → δ) : δ := data.foldl d (λ b d, b.foldl (λ r a, f r a.1 a.2) d) -lemma foldl_eq_lem {γ : Type u} {δ : Type w} (d : δ) (f : δ → γ → δ) : Π l : list (list γ), +theorem foldl_eq_lem {γ : Type u} {δ : Type w} (d : δ) (f : δ → γ → δ) : Π l : list (list γ), l.foldr (λ (b:list γ) d, b.foldl f d) d = (l.foldr (λ(bkt r : list γ), r ++ bkt) []).foldl f d | [] := rfl | (l::ls) := show l.foldl f (ls.foldr (λ (b:list γ) d, b.foldl f d) d) = @@ -142,7 +142,7 @@ theorem valid_aux.eq {idx : α → nat} : Π {l : list (list Σ a, β a)} {sz : have idx a = list.length l - 1 - i, from valid_aux.eq v el, by rwa [nat.sub_sub, nat.add_comm] at this -theorem valid_aux.insert_lemma1 {idx : α → nat} : Π {l : list (list Σ a, β a)} {sz : nat}, valid_aux idx l sz → +private lemma valid_aux.insert_lemma1 {idx : α → nat} : Π {l : list (list Σ a, β a)} {sz : nat}, valid_aux idx l sz → ∀ {i h a b}, sigma.mk a b ∈ l.nth_le i h → idx a = l.length - 1 - i | ._ ._ valid_aux.nil i h _ _ _ := absurd h (nat.not_lt_zero _) | ._ ._ (@valid_aux.cons ._ ._ ._ c l sz v nd e) 0 h a b el := e ⟨a, b⟩ el @@ -238,7 +238,7 @@ section local notation `L` := array.read bkts bidx private def bkts' : bucket_array α β n := array.write bkts bidx (f L) - theorem valid.modify_aux1 {δ fn} {b : δ} : Π (i) (h : i ≤ n.1) (hb : i ≤ bidx.1), + private lemma valid.modify_aux1 {δ fn} {b : δ} : Π (i) (h : i ≤ n.1) (hb : i ≤ bidx.1), array.iterate_aux bkts fn i h b = array.iterate_aux bkts' fn i h b | 0 h hb := by simp[array.iterate_aux] | (i+1) h (hb : i < bidx.1) := by simp[array.iterate_aux]; exact @@ -250,7 +250,7 @@ section (hfl : f L = u ++ v2 ++ w) include hl hfl - theorem append_of_modify_aux : Π (i) (h : i ≤ n.1) (hb : i > bidx.1), + private lemma append_of_modify_aux : Π (i) (h : i ≤ n.1) (hb : i > bidx.1), ∃ u' w', array.iterate_aux bkts (λ_ bkt r, r ++ bkt) i h [] = u' ++ v1 ++ w' ∧ array.iterate_aux bkts' (λ_ bkt r, r ++ bkt) i h [] = u' ++ v2 ++ w' | 0 _ hb := absurd hb (nat.not_lt_zero _) @@ -287,7 +287,7 @@ section (djwv : (w.map sigma.fst).disjoint (v2.map sigma.fst)) include hvnd hal djuv djwv - theorem valid.modify_aux2 : Π (i) (h : i ≤ n.1) (hb : i > bidx.1) {sz : ℕ}, + private lemma valid.modify_aux2 : Π (i) (h : i ≤ n.1) (hb : i > bidx.1) {sz : ℕ}, valid_aux (λa, (mk_idx n (hash_fn a)).1) (array.iterate_aux bkts (λ_ v l, v :: l) i h []) sz → sz + v2.length ≥ v1.length ∧ valid_aux (λa, (mk_idx n (hash_fn a)).1) (array.iterate_aux bkts' (λ_ v l, v :: l) i h []) (sz + v2.length - v1.length) | 0 _ hb sz := absurd hb (nat.not_lt_zero _) @@ -495,7 +495,7 @@ theorem not_contains_empty (hash_fn : α → nat) (n a) : ¬ (@mk_hash_map α _ β hash_fn n).contains a := by apply bool_iff_false.2; dsimp [contains]; rw [find_empty]; refl -lemma insert_lemma (hash_fn : α → nat) {n n'} +theorem insert_theorem (hash_fn : α → nat) {n n'} {bkts : bucket_array α β n} {sz} (v : valid hash_fn bkts sz) : valid hash_fn (bkts.foldl (mk_array _ [] : bucket_array α β n') (reinsert_aux hash_fn)) sz := suffices ∀ (l : list Σ a, β a), @@ -568,7 +568,7 @@ let n' : ℕ+ := ⟨n.1 * 2, mul_pos n.2 dec_trivial⟩, size := size', nbuckets := n', buckets := buckets'', - is_valid := insert_lemma _ valid' } + is_valid := insert_theorem _ valid' } theorem mem_insert : Π (m : hash_map α β) (a b a' b'), sigma.mk a' b' ∈ (m.insert a b).entries ↔ diff --git a/data/int/basic.lean b/data/int/basic.lean index f702b1f414d13..24f186d4cb209 100644 --- a/data/int/basic.lean +++ b/data/int/basic.lean @@ -27,7 +27,7 @@ begin simp [*, sub_nat_nat] end -@[simp] lemma neg_add_neg (m n : ℕ) : -[1+m] + -[1+n] = -[1+nat.succ(m+n)] := rfl +@[simp] theorem neg_add_neg (m n : ℕ) : -[1+m] + -[1+n] = -[1+nat.succ(m+n)] := rfl end int diff --git a/data/list/basic.lean b/data/list/basic.lean index 346fca2118f83..0b4f376fc4de8 100644 --- a/data/list/basic.lean +++ b/data/list/basic.lean @@ -15,18 +15,18 @@ variables {α : Type u} {β : Type v} /- theorems -/ @[simp] -lemma cons_ne_nil (a : α) (l : list α) : a::l ≠ [] := +theorem cons_ne_nil (a : α) (l : list α) : a::l ≠ [] := begin intro, contradiction end -lemma head_eq_of_cons_eq {α : Type} {h₁ h₂ : α} {t₁ t₂ : list α} : +theorem head_eq_of_cons_eq {α : Type} {h₁ h₂ : α} {t₁ t₂ : list α} : (h₁::t₁) = (h₂::t₂) → h₁ = h₂ := assume Peq, list.no_confusion Peq (assume Pheq Pteq, Pheq) -lemma tail_eq_of_cons_eq {α : Type} {h₁ h₂ : α} {t₁ t₂ : list α} : +theorem tail_eq_of_cons_eq {α : Type} {h₁ h₂ : α} {t₁ t₂ : list α} : (h₁::t₁) = (h₂::t₂) → t₁ = t₂ := assume Peq, list.no_confusion Peq (assume Pheq Pteq, Pteq) -lemma cons_inj {α : Type} {a : α} : injective (cons a) := +theorem cons_inj {α : Type} {a : α} : injective (cons a) := assume l₁ l₂, assume Pe, tail_eq_of_cons_eq Pe /- append -/ @@ -68,11 +68,11 @@ by induction l₂ with b l₂ ih; simp /- last -/ @[simp] -lemma last_singleton (a : α) (h : [a] ≠ []) : last [a] h = a := +theorem last_singleton (a : α) (h : [a] ≠ []) : last [a] h = a := rfl @[simp] -lemma last_cons_cons (a₁ a₂ : α) (l : list α) (h : a₁::a₂::l ≠ []) : +theorem last_cons_cons (a₁ a₂ : α) (l : list α) (h : a₁::a₂::l ≠ []) : last (a₁::a₂::l) h = last (a₂::l) (cons_ne_nil a₂ l) := rfl @@ -130,7 +130,7 @@ list.rec_on l show index_of a (b::l) = length (b::l), begin rw [index_of_cons, if_neg this.left, ih this.right], reflexivity end) -lemma index_of_le_length {a : α} {l : list α} : index_of a l ≤ length l := +theorem index_of_le_length {a : α} {l : list α} : index_of a l ≤ length l := list.rec_on l (by simp) (assume b l, assume ih : index_of a l ≤ length l, @@ -139,7 +139,7 @@ list.rec_on l (assume : a = b, begin simp [this, index_of_cons_of_eq l (eq.refl b)], apply zero_le end) (assume : a ≠ b, begin rw [index_of_cons_of_ne l this], apply succ_le_succ ih end)) -lemma not_mem_of_index_of_eq_length : ∀ {a : α} {l : list α}, index_of a l = length l → a ∉ l +theorem not_mem_of_index_of_eq_length : ∀ {a : α} {l : list α}, index_of a l = length l → a ∉ l | a [] := by simp | a (b::l) := begin @@ -151,7 +151,7 @@ lemma not_mem_of_index_of_eq_length : ∀ {a : α} {l : list α}, index_of a l = exact not_mem_of_index_of_eq_length (nat.succ_inj h) end -lemma index_of_lt_length {a} {l : list α} (al : a ∈ l) : index_of a l < length l := +theorem index_of_lt_length {a} {l : list α} (al : a ∈ l) : index_of a l < length l := begin apply lt_of_le_of_ne, apply index_of_le_length, @@ -207,11 +207,11 @@ def ith : Π (l : list α) (i : nat), i < length l → α | (a::ains) (succ i) h := ith ains i (lt_of_succ_lt_succ h) @[simp] -lemma ith_zero (a : α) (l : list α) (h : 0 < length (a::l)) : ith (a::l) 0 h = a := +theorem ith_zero (a : α) (l : list α) (h : 0 < length (a::l)) : ith (a::l) 0 h = a := rfl @[simp] -lemma ith_succ (a : α) (l : list α) (i : nat) (h : succ i < length (a::l)) +theorem ith_succ (a : α) (l : list α) (i : nat) (h : succ i < length (a::l)) : ith (a::l) (succ i) h = ith l i (lt_of_succ_lt_succ h) := rfl end ith @@ -219,22 +219,22 @@ end ith section take @[simp] -lemma taken_zero : ∀ (l : list α), take 0 l = [] := +theorem taken_zero : ∀ (l : list α), take 0 l = [] := begin intros, reflexivity end @[simp] -lemma taken_nil : ∀ n, take n [] = ([] : list α) +theorem taken_nil : ∀ n, take n [] = ([] : list α) | 0 := rfl | (n+1) := rfl -lemma taken_cons : ∀ n (a : α) (l : list α), take (succ n) (a::l) = a :: take n l := +theorem taken_cons : ∀ n (a : α) (l : list α), take (succ n) (a::l) = a :: take n l := begin intros, reflexivity end -lemma taken_all : ∀ (l : list α), take (length l) l = l +theorem taken_all : ∀ (l : list α), take (length l) l = l | [] := rfl | (a::l) := begin change a :: (take (length l) l) = a :: l, rw taken_all end -lemma taken_all_of_ge : ∀ {n} {l : list α}, n ≥ length l → take n l = l +theorem taken_all_of_ge : ∀ {n} {l : list α}, n ≥ length l → take n l = l | 0 [] h := rfl | 0 (a::l) h := absurd h (not_le_of_gt (zero_lt_succ _)) | (n+1) [] h := rfl @@ -246,14 +246,14 @@ lemma taken_all_of_ge : ∀ {n} {l : list α}, n ≥ length l → take n l = l -- TODO(Jeremy): restore when we have min /- -lemma taken_taken : ∀ (n m) (l : list α), take n (take m l) = take (min n m) l +theorem taken_taken : ∀ (n m) (l : list α), take n (take m l) = take (min n m) l | n 0 l := sorry -- by rewrite [min_zero, taken_zero, taken_nil] | 0 m l := sorry -- by rewrite [zero_min] | (succ n) (succ m) nil := sorry -- by rewrite [*taken_nil] | (succ n) (succ m) (a::l) := sorry -- by rewrite [*taken_cons, taken_taken, min_succ_succ] -/ -lemma length_taken_le : ∀ (n) (l : list α), length (take n l) ≤ n +theorem length_taken_le : ∀ (n) (l : list α), length (take n l) ≤ n | 0 l := begin rw [taken_zero], reflexivity end | (succ n) (a::l) := begin rw [taken_cons, length_cons], apply succ_le_succ, @@ -263,7 +263,7 @@ lemma length_taken_le : ∀ (n) (l : list α), length (take n l) ≤ n -- TODO(Jeremy): restore when we have min /- -lemma length_taken_eq : ∀ (n) (l : list α), length (take n l) = min n (length l) +theorem length_taken_eq : ∀ (n) (l : list α), length (take n l) = min n (length l) | 0 l := sorry -- by rewrite [taken_zero, zero_min] | (succ n) (a::l) := sorry -- by rewrite [taken_cons, *length_cons, *add_one, min_succ_succ, length_taken_eq] @@ -297,14 +297,14 @@ def count (a : α) : list α → nat | (x::xs) := if a = x then succ (count xs) else count xs @[simp] -lemma count_nil (a : α) : count a [] = 0 := +theorem count_nil (a : α) : count a [] = 0 := rfl -lemma count_cons (a b : α) (l : list α) : +theorem count_cons (a b : α) (l : list α) : count a (b :: l) = if a = b then succ (count a l) else count a l := rfl -lemma count_cons' (a b : α) (l : list α) : +theorem count_cons' (a b : α) (l : list α) : count a (b :: l) = count a l + (if a = b then 1 else 0) := decidable.by_cases (assume : a = b, begin rw [count_cons, if_pos this, if_pos this] end) @@ -312,14 +312,14 @@ decidable.by_cases @[simp] -lemma count_cons_self (a : α) (l : list α) : count a (a::l) = succ (count a l) := +theorem count_cons_self (a : α) (l : list α) : count a (a::l) = succ (count a l) := if_pos rfl @[simp] -lemma count_cons_of_ne {a b : α} (h : a ≠ b) (l : list α) : count a (b::l) = count a l := +theorem count_cons_of_ne {a b : α} (h : a ≠ b) (l : list α) : count a (b::l) = count a l := if_neg h -lemma count_cons_ge_count (a b : α) (l : list α) : count a (b :: l) ≥ count a l := +theorem count_cons_ge_count (a b : α) (l : list α) : count a (b :: l) ≥ count a l := decidable.by_cases (assume : a = b, begin subst b, rewrite count_cons_self, apply le_succ end) (assume : a ≠ b, begin rw (count_cons_of_ne this), apply le_refl end) @@ -327,11 +327,11 @@ decidable.by_cases -- TODO(Jeremy): without the reflexivity, this yields the goal "1 = 1". the first is from has_one, -- the second is succ 0. Make sure the simplifier can eventually handle this. -lemma count_singleton (a : α) : count a [a] = 1 := +theorem count_singleton (a : α) : count a [a] = 1 := by simp @[simp] -lemma count_append (a : α) : ∀ l₁ l₂, count a (l₁ ++ l₂) = count a l₁ + count a l₂ +theorem count_append (a : α) : ∀ l₁ l₂, count a (l₁ ++ l₂) = count a l₁ + count a l₂ | [] l₂ := begin rw [nil_append, count_nil, zero_add] end | (b::l₁) l₂ := decidable.by_cases (assume : a = b, by rw [←this, cons_append, count_cons_self, count_cons_self, succ_add, @@ -339,10 +339,10 @@ lemma count_append (a : α) : ∀ l₁ l₂, count a (l₁ ++ l₂) = count a l (assume : a ≠ b, by rw [cons_append, count_cons_of_ne this, count_cons_of_ne this, count_append]) @[simp] -lemma count_concat (a : α) (l : list α) : count a (concat l a) = succ (count a l) := +theorem count_concat (a : α) (l : list α) : count a (concat l a) = succ (count a l) := by rw [concat_eq_append, count_append, count_singleton] -lemma mem_of_count_pos : ∀ {a : α} {l : list α}, count a l > 0 → a ∈ l +theorem mem_of_count_pos : ∀ {a : α} {l : list α}, count a l > 0 → a ∈ l | a [] h := absurd h (lt_irrefl _) | a (b::l) h := decidable.by_cases (assume : a = b, begin subst b, apply mem_cons_self end) @@ -351,7 +351,7 @@ lemma mem_of_count_pos : ∀ {a : α} {l : list α}, count a l > 0 → a ∈ l have a ∈ l, from mem_of_count_pos this, show a ∈ b::l, from mem_cons_of_mem _ this) -lemma count_pos_of_mem : ∀ {a : α} {l : list α}, a ∈ l → count a l > 0 +theorem count_pos_of_mem : ∀ {a : α} {l : list α}, a ∈ l → count a l > 0 | a [] h := absurd h (not_mem_nil _) | a (b::l) h := or.elim h (assume : a = b, begin subst b, rw count_cons_self, apply zero_lt_succ end) @@ -359,11 +359,11 @@ lemma count_pos_of_mem : ∀ {a : α} {l : list α}, a ∈ l → count a l > 0 count a (b::l) ≥ count a l : count_cons_ge_count _ _ _ ... > 0 : count_pos_of_mem this) -lemma mem_iff_count_pos (a : α) (l : list α) : a ∈ l ↔ count a l > 0 := +theorem mem_iff_count_pos (a : α) (l : list α) : a ∈ l ↔ count a l > 0 := iff.intro count_pos_of_mem mem_of_count_pos @[simp] -lemma count_eq_zero_of_not_mem {a : α} {l : list α} (h : a ∉ l) : count a l = 0 := +theorem count_eq_zero_of_not_mem {a : α} {l : list α} (h : a ∉ l) : count a l = 0 := have ∀ n, count a l = n → count a l = 0, begin intro n, cases n, @@ -372,7 +372,7 @@ have ∀ n, count a l = n → count a l = 0, end, this (count a l) rfl -lemma not_mem_of_count_eq_zero {a : α} {l : list α} (h : count a l = 0) : a ∉ l := +theorem not_mem_of_count_eq_zero {a : α} {l : list α} (h : count a l = 0) : a ∉ l := assume : a ∈ l, have count a l > 0, from count_pos_of_mem this, show false, begin rw h at this, exact nat.not_lt_zero _ this end diff --git a/data/list/comb.lean b/data/list/comb.lean index 6dfbc7f7b6c48..4f1c29d6e6019 100644 --- a/data/list/comb.lean +++ b/data/list/comb.lean @@ -477,13 +477,13 @@ variable [h : decidable_pred p] include h variable {f : Π a, p a → β} -lemma dmap_nil : dmap p f [] = [] := rfl -lemma dmap_cons_of_pos {a : α} (P : p a) : ∀ l, dmap p f (a::l) = (f a P) :: dmap p f l := +theorem dmap_nil : dmap p f [] = [] := rfl +theorem dmap_cons_of_pos {a : α} (P : p a) : ∀ l, dmap p f (a::l) = (f a P) :: dmap p f l := λ l, dif_pos P -lemma dmap_cons_of_neg {a : α} (P : ¬ p a) : ∀ l, dmap p f (a::l) = dmap p f l := +theorem dmap_cons_of_neg {a : α} (P : ¬ p a) : ∀ l, dmap p f (a::l) = dmap p f l := λ l, dif_neg P -lemma mem_dmap : ∀ {l : list α} {a} (Pa : p a), a ∈ l → (f a Pa) ∈ dmap p f l +theorem mem_dmap : ∀ {l : list α} {a} (Pa : p a), a ∈ l → (f a Pa) ∈ dmap p f l | [] := assume a Pa Pinnil, absurd Pinnil (not_mem_nil _) | (a::l) := assume b Pb Pbin, or.elim (eq_or_mem_of_mem_cons Pbin) (assume Pbeqa, begin @@ -503,7 +503,7 @@ lemma mem_dmap : ∀ {l : list α} {a} (Pa : p a), a ∈ l → (f a Pa) ∈ dmap exact mem_dmap Pb Pbinl end) -lemma exists_of_mem_dmap : ∀ {l : list α} {b : β}, b ∈ dmap p f l → ∃ a P, a ∈ l ∧ b = f a P +theorem exists_of_mem_dmap : ∀ {l : list α} {b : β}, b ∈ dmap p f l → ∃ a P, a ∈ l ∧ b = f a P | [] := assume b, begin rw dmap_nil, intro h, exact absurd h (not_mem_nil _) end | (a::l) := assume b, if Pa : p a then @@ -526,7 +526,7 @@ lemma exists_of_mem_dmap : ∀ {l : list α} {b : β}, b ∈ dmap p f l → ∃ (and.right P'))) end -lemma map_dmap_of_inv_of_pos {g : β → α} (Pinv : ∀ a (Pa : p a), g (f a Pa) = a) : +theorem map_dmap_of_inv_of_pos {g : β → α} (Pinv : ∀ a (Pa : p a), g (f a Pa) = a) : ∀ {l : list α}, (∀ ⦃a⦄, a ∈ l → p a) → map g (dmap p f l) = l | [] := assume Pl, by rw [dmap_nil]; reflexivity | (a::l) := assume Pal, @@ -535,7 +535,7 @@ lemma map_dmap_of_inv_of_pos {g : β → α} (Pinv : ∀ a (Pa : p a), g (f a Pa from assume x Pxin, Pal (mem_cons_of_mem a Pxin), by rw [dmap_cons_of_pos Pa, map_cons, Pinv, map_dmap_of_inv_of_pos Pl] -lemma mem_of_dinj_of_mem_dmap (Pdi : dinj p f) : +theorem mem_of_dinj_of_mem_dmap (Pdi : dinj p f) : ∀ {l : list α} {a} (Pa : p a), (f a Pa) ∈ dmap p f l → a ∈ l | [] := assume a Pa Pinnil, absurd Pinnil (not_mem_nil _) | (b::l) := assume a Pa Pmap, @@ -555,7 +555,7 @@ lemma mem_of_dinj_of_mem_dmap (Pdi : dinj p f) : exact mem_of_dinj_of_mem_dmap Pa Pmap end -lemma not_mem_dmap_of_dinj_of_not_mem (Pdi : dinj p f) {l : list α} {a} (Pa : p a) : +theorem not_mem_dmap_of_dinj_of_not_mem (Pdi : dinj p f) {l : list α} {a} (Pa : p a) : a ∉ l → (f a Pa) ∉ dmap p f l := mt (mem_of_dinj_of_mem_dmap Pdi Pa) diff --git a/data/list/perm.lean b/data/list/perm.lean index 69a6b6160064e..01522cd6442ff 100644 --- a/data/list/perm.lean +++ b/data/list/perm.lean @@ -310,7 +310,7 @@ end count -- Auxiliary theorem for performing cases-analysis on l₂. -- We use it to prove perm_inv_core. -private theorem discr {P : Prop} {a b : α} {l₁ l₂ l₃ : list α} : +private lemma discr {P : Prop} {a b : α} {l₁ l₂ l₃ : list α} : a::l₁ = l₂++(b::l₃) → (l₂ = [] → a = b → l₁ = l₃ → P) → (∀ t, l₂ = a::t → l₁ = t++(b::l₃) → P) → P := @@ -327,7 +327,7 @@ end -- Auxiliary theorem for performing cases-analysis on l₂. -- We use it to prove perm_inv_core. -private theorem discr₂ {P : Prop} {a b c : α} {l₁ l₂ l₃ : list α} +private lemma discr₂ {P : Prop} {a b c : α} {l₁ l₂ l₃ : list α} (e : a::b::l₁ = l₂++(c::l₃)) (H₁ : l₂ = [] → l₃ = b::l₁ → a = c → P) (H₂ : l₂ = [a] → b = c → l₁ = l₃ → P) @@ -348,7 +348,7 @@ open qeq notation l' `≈`:50 a `|` l:50 := qeq a l l' -lemma perm_of_qeq {a : α} {l₁ l₂ : list α} : l₁≈a|l₂ → l₁~a::l₂ := +theorem perm_of_qeq {a : α} {l₁ l₂ : list α} : l₁≈a|l₂ → l₁~a::l₂ := assume q, qeq.rec_on q (λ h, perm.refl (a :: h)) (λ b t₁ t₂ q₁ r₁, calc @@ -622,7 +622,7 @@ begin begin simp [ha₁, ha₂], apply ih, apply perm_app_left, exact h end end -lemma perm_insert_insert (x y : α) (l : list α) : +theorem perm_insert_insert (x y : α) (l : list α) : insert x (insert y l) ~ insert y (insert x l) := if yl : y ∈ l then if xl : x ∈ l then by simp [xl, yl] diff --git a/data/list/set.lean b/data/list/set.lean index 0539b60e2e4c7..e903b686df069 100644 --- a/data/list/set.lean +++ b/data/list/set.lean @@ -71,22 +71,22 @@ section erase variable [decidable_eq α] @[simp] -lemma erase_nil (a : α) : [].erase a = [] := +theorem erase_nil (a : α) : [].erase a = [] := rfl -lemma erase_cons (a b : α) (l : list α) : (b :: l).erase a = if b = a then l else b :: l.erase a := +theorem erase_cons (a b : α) (l : list α) : (b :: l).erase a = if b = a then l else b :: l.erase a := rfl @[simp] -lemma erase_cons_head (a : α) (l : list α) : (a :: l).erase a = l := +theorem erase_cons_head (a : α) (l : list α) : (a :: l).erase a = l := by simp [erase_cons, if_pos] @[simp] -lemma erase_cons_tail {a b : α} (l : list α) (h : b ≠ a) : (b::l).erase a = b :: l.erase a := +theorem erase_cons_tail {a b : α} (l : list α) (h : b ≠ a) : (b::l).erase a = b :: l.erase a := by simp [erase_cons, if_neg, h] @[simp] -lemma length_erase_of_mem {a : α} : ∀{l:list α}, a ∈ l → length (l.erase a) = pred (length l) +theorem length_erase_of_mem {a : α} : ∀{l:list α}, a ∈ l → length (l.erase a) = pred (length l) | [] h := rfl | [x] h := begin simp at h, simp [h] end | (x::y::xs) h := if h' : x = a then @@ -96,35 +96,35 @@ lemma length_erase_of_mem {a : α} : ∀{l:list α}, a ∈ l → length (l.erase by simp [h', length_erase_of_mem ainyxs, one_add] @[simp] -lemma erase_of_not_mem {a : α} : ∀{l : list α}, a ∉ l → l.erase a = l +theorem erase_of_not_mem {a : α} : ∀{l : list α}, a ∉ l → l.erase a = l | [] h := rfl | (x::xs) h := have anex : x ≠ a, from λ aeqx : x = a, absurd (or.inl aeqx.symm) h, have aninxs : a ∉ xs, from λ ainxs : a ∈ xs, absurd (or.inr ainxs) h, by simp [anex, erase_of_not_mem aninxs] -lemma erase_append_left {a : α} : ∀ {l₁:list α} (l₂), a ∈ l₁ → (l₁++l₂).erase a = l₁.erase a ++ l₂ +theorem erase_append_left {a : α} : ∀ {l₁:list α} (l₂), a ∈ l₁ → (l₁++l₂).erase a = l₁.erase a ++ l₂ | [] l₂ h := absurd h (not_mem_nil a) | (x::xs) l₂ h := if h' : x = a then by simp [h'] else have a ∈ xs, from mem_of_ne_of_mem (assume h, h' h.symm) h, by simp [erase_append_left l₂ this, h'] -lemma erase_append_right {a : α} : ∀{l₁ : list α} (l₂), a ∉ l₁ → (l₁++l₂).erase a = l₁ ++ l₂.erase a +theorem erase_append_right {a : α} : ∀{l₁ : list α} (l₂), a ∉ l₁ → (l₁++l₂).erase a = l₁ ++ l₂.erase a | [] l₂ h := rfl | (x::xs) l₂ h := if h' : x = a then begin simp [h'] at h, contradiction end else have a ∉ xs, from not_mem_of_not_mem_cons h, by simp [erase_append_right l₂ this, h'] -lemma erase_sublist (a : α) : ∀(l : list α), l.erase a <+ l +theorem erase_sublist (a : α) : ∀(l : list α), l.erase a <+ l | [] := sublist.refl nil | (x :: xs) := if h : x = a then by simp [h] else begin simp [h], apply cons_sublist_cons, apply erase_sublist xs end -lemma erase_subset (a : α) (l : list α) : l.erase a ⊆ l := +theorem erase_subset (a : α) (l : list α) : l.erase a ⊆ l := subset_of_sublist (erase_sublist a l) theorem mem_erase_of_ne_of_mem {a b : α} : ∀ {l : list α}, a ≠ b → a ∈ l → a ∈ l.erase b @@ -155,54 +155,54 @@ section disjoint def disjoint (l₁ l₂ : list α) : Prop := ∀ ⦃a⦄, (a ∈ l₁ → a ∈ l₂ → false) -lemma disjoint_left {l₁ l₂ : list α} : disjoint l₁ l₂ → ∀ {a}, a ∈ l₁ → a ∉ l₂ := +theorem disjoint_left {l₁ l₂ : list α} : disjoint l₁ l₂ → ∀ {a}, a ∈ l₁ → a ∉ l₂ := λ d, d -lemma disjoint_right {l₁ l₂ : list α} : disjoint l₁ l₂ → ∀ {a}, a ∈ l₂ → a ∉ l₁ := +theorem disjoint_right {l₁ l₂ : list α} : disjoint l₁ l₂ → ∀ {a}, a ∈ l₂ → a ∉ l₁ := λ d a i₂ i₁, d i₁ i₂ -lemma disjoint.comm {l₁ l₂ : list α} : disjoint l₁ l₂ → disjoint l₂ l₁ := +theorem disjoint.comm {l₁ l₂ : list α} : disjoint l₁ l₂ → disjoint l₂ l₁ := λ d a i₂ i₁, d i₁ i₂ -lemma disjoint_of_subset_left {l₁ l₂ l : list α} : l₁ ⊆ l → disjoint l l₂ → disjoint l₁ l₂ := +theorem disjoint_of_subset_left {l₁ l₂ l : list α} : l₁ ⊆ l → disjoint l l₂ → disjoint l₁ l₂ := λ ss d x xinl₁, d (ss xinl₁) -lemma disjoint_of_subset_right {l₁ l₂ l : list α} : l₂ ⊆ l → disjoint l₁ l → disjoint l₁ l₂ := +theorem disjoint_of_subset_right {l₁ l₂ l : list α} : l₂ ⊆ l → disjoint l₁ l → disjoint l₁ l₂ := λ ss d x xinl xinl₁, d xinl (ss xinl₁) -lemma disjoint_of_disjoint_cons_left {a : α} {l₁ l₂} : disjoint (a::l₁) l₂ → disjoint l₁ l₂ := +theorem disjoint_of_disjoint_cons_left {a : α} {l₁ l₂} : disjoint (a::l₁) l₂ → disjoint l₁ l₂ := disjoint_of_subset_left (list.subset_cons _ _) -lemma disjoint_of_disjoint_cons_right {a : α} {l₁ l₂} : disjoint l₁ (a::l₂) → disjoint l₁ l₂ := +theorem disjoint_of_disjoint_cons_right {a : α} {l₁ l₂} : disjoint l₁ (a::l₂) → disjoint l₁ l₂ := disjoint_of_subset_right (list.subset_cons _ _) -lemma disjoint_nil_left (l : list α) : disjoint [] l := +theorem disjoint_nil_left (l : list α) : disjoint [] l := λ a ab, absurd ab (not_mem_nil a) -lemma disjoint_nil_right (l : list α) : disjoint l [] := +theorem disjoint_nil_right (l : list α) : disjoint l [] := disjoint.comm (disjoint_nil_left l) -lemma disjoint_cons_of_not_mem_of_disjoint {a : α} {l₁ l₂ : list α} : +theorem disjoint_cons_of_not_mem_of_disjoint {a : α} {l₁ l₂ : list α} : a ∉ l₂ → disjoint l₁ l₂ → disjoint (a::l₁) l₂ := λ nainl₂ d x (xinal₁ : x ∈ a::l₁), or.elim (eq_or_mem_of_mem_cons xinal₁) (λ xeqa : x = a, eq.symm xeqa ▸ nainl₂) (λ xinl₁ : x ∈ l₁, disjoint_left d xinl₁) -lemma disjoint_append_of_disjoint_left {l₁ l₂ l : list α} : +theorem disjoint_append_of_disjoint_left {l₁ l₂ l : list α} : disjoint l₁ l → disjoint l₂ l → disjoint (l₁++l₂) l := λ d₁ d₂ x h, or.elim (mem_or_mem_of_mem_append h) (@d₁ x) (@d₂ x) -lemma disjoint_of_disjoint_append_left_left {l₁ l₂ l : list α} : disjoint (l₁++l₂) l → disjoint l₁ l := +theorem disjoint_of_disjoint_append_left_left {l₁ l₂ l : list α} : disjoint (l₁++l₂) l → disjoint l₁ l := disjoint_of_subset_left (list.subset_append_left _ _) -lemma disjoint_of_disjoint_append_left_right {l₁ l₂ l : list α} : disjoint (l₁++l₂) l → disjoint l₂ l := +theorem disjoint_of_disjoint_append_left_right {l₁ l₂ l : list α} : disjoint (l₁++l₂) l → disjoint l₂ l := disjoint_of_subset_left (list.subset_append_right _ _) -lemma disjoint_of_disjoint_append_right_left {l₁ l₂ l : list α} : disjoint l (l₁++l₂) → disjoint l l₁ := +theorem disjoint_of_disjoint_append_right_left {l₁ l₂ l : list α} : disjoint l (l₁++l₂) → disjoint l l₁ := disjoint_of_subset_right (list.subset_append_left _ _) -lemma disjoint_of_disjoint_append_right_right {l₁ l₂ l : list α} : disjoint l (l₁++l₂) → disjoint l l₂ := +theorem disjoint_of_disjoint_append_right_right {l₁ l₂ l : list α} : disjoint l (l₁++l₂) → disjoint l l₂ := disjoint_of_subset_right (list.subset_append_right _ _) end disjoint @@ -251,7 +251,7 @@ begin simp [ieq] end -lemma upto_step : ∀ (n : nat), upto (succ n) = (map succ (upto n)) ++ [0] +theorem upto_step : ∀ (n : nat), upto (succ n) = (map succ (upto n)) ++ [0] | 0 := rfl | (succ n) := by simp [(upto_step n).symm] @@ -683,7 +683,7 @@ theorem nodup_filter (p : α → Prop) [decidable_pred p] : (λ pa : p a, begin rw [filter_cons_of_pos _ pa], exact (nodup_cons nainf ndf) end) (λ npa : ¬ p a, begin rw [filter_cons_of_neg _ npa], exact ndf end) -lemma dmap_nodup_of_dinj {p : α → Prop} [h : decidable_pred p] {f : Π a, p a → β} (pdi : dinj p f) : +theorem dmap_nodup_of_dinj {p : α → Prop} [h : decidable_pred p] {f : Π a, p a → β} (pdi : dinj p f) : ∀ {l : list α}, nodup l → nodup (dmap p f l) | [] := assume P, nodup.ndnil | (a::l) := assume Pnodup, diff --git a/data/list/sort.lean b/data/list/sort.lean index 4704fd4e0d2b8..3b93fd20c350a 100644 --- a/data/list/sort.lean +++ b/data/list/sort.lean @@ -34,10 +34,10 @@ iff.intro apply add_pos_right _ npos end -lemma succ_le_succ_iff (m n : ℕ) : succ m ≤ succ n ↔ m ≤ n := +theorem succ_le_succ_iff (m n : ℕ) : succ m ≤ succ n ↔ m ≤ n := ⟨le_of_succ_le_succ, succ_le_succ⟩ -lemma lt_succ_iff_le (m n : ℕ) : m < succ n ↔ m ≤ n := +theorem lt_succ_iff_le (m n : ℕ) : m < succ n ↔ m ≤ n := succ_le_succ_iff m n end nat diff --git a/data/nat/basic.lean b/data/nat/basic.lean index 9c071b77f6dee..4175af3b16137 100644 --- a/data/nat/basic.lean +++ b/data/nat/basic.lean @@ -20,10 +20,10 @@ def addl : ℕ → ℕ → ℕ | 0 y := y local infix ` ⊕ `:65 := addl -@[simp] lemma addl_zero_left (n : ℕ) : 0 ⊕ n = n := rfl -@[simp] lemma addl_succ_left (m n : ℕ) : succ m ⊕ n = succ (m ⊕ n) := rfl +@[simp] theorem addl_zero_left (n : ℕ) : 0 ⊕ n = n := rfl +@[simp] theorem addl_succ_left (m n : ℕ) : succ m ⊕ n = succ (m ⊕ n) := rfl -@[simp] lemma zero_has_zero : nat.zero = 0 := rfl +@[simp] theorem zero_has_zero : nat.zero = 0 := rfl local attribute [simp] nat.add_zero nat.add_succ nat.zero_add nat.succ_add @@ -53,7 +53,7 @@ nat.succ.inj_arrow H id theorem discriminate {B : Type _} {n : ℕ} (H1: n = 0 → B) (H2 : ∀m, n = succ m → B) : B := by ginduction n with h; [exact H1 h, exact H2 _ h] -lemma one_succ_zero : 1 = succ 0 := rfl +theorem one_succ_zero : 1 = succ 0 := rfl --local attribute [simp] one_succ_zero theorem two_step_induction {P : ℕ → Sort u} (H1 : P 0) (H2 : P 1) diff --git a/data/nat/gcd.lean b/data/nat/gcd.lean index 45631d38a1cf0..dfbd07551cfeb 100644 --- a/data/nat/gcd.lean +++ b/data/nat/gcd.lean @@ -136,7 +136,7 @@ dvd_antisymm /- coprime -/ -lemma gcd_eq_one_of_coprime {m n : ℕ} : coprime m n → gcd m n = 1 := +theorem gcd_eq_one_of_coprime {m n : ℕ} : coprime m n → gcd m n = 1 := λ h, h theorem coprime_swap {m n : ℕ} (H : coprime n m) : coprime m n := diff --git a/data/nat/sub.lean b/data/nat/sub.lean index 6331466a79843..a0a8950ecf52e 100644 --- a/data/nat/sub.lean +++ b/data/nat/sub.lean @@ -137,7 +137,7 @@ theorem dist_mul_left (k n m : ℕ) : dist (k * n) (k * m) = k * dist n m := by rw [mul_comm k n, mul_comm k m, dist_mul_right, mul_comm] -- TODO(Jeremy): do when we have max and minx ---lemma dist_eq_max_sub_min {i j : nat} : dist i j = (max i j) - min i j := +--theorem dist_eq_max_sub_min {i j : nat} : dist i j = (max i j) - min i j := --sorry /- or.elim (lt_or_ge i j) @@ -147,10 +147,10 @@ or.elim (lt_or_ge i j) by rw [max_eq_left this , min_eq_right this, dist_eq_sub_of_ge this]) -/ -lemma dist_succ_succ {i j : nat} : dist (succ i) (succ j) = dist i j := +theorem dist_succ_succ {i j : nat} : dist (succ i) (succ j) = dist i j := by simp [dist.def, succ_sub_succ] -lemma dist_pos_of_ne {i j : nat} : i ≠ j → dist i j > 0 := +theorem dist_pos_of_ne {i j : nat} : i ≠ j → dist i j > 0 := assume hne, nat.lt_by_cases (assume : i < j, begin rw [dist_eq_sub_of_le (le_of_lt this)], apply nat.sub_pos_of_lt this end) diff --git a/data/num/lemmas.lean b/data/num/lemmas.lean index 91a5dea6a0347..e6ad0bc2b263c 100644 --- a/data/num/lemmas.lean +++ b/data/num/lemmas.lean @@ -164,7 +164,7 @@ namespace pos_num by induction m with m IH m IH; intro n; cases n with n n; try {unfold cmp}; try {refl}; rw ←IH; cases cmp m n; refl - lemma cmp_dec_lemma {m n} : m < n → bit1 m < bit0 n := + theorem cmp_dec_theorem {m n} : m < n → bit1 m < bit0 n := show (m:ℕ) < n → (m + m + 1 + 1 : ℕ) ≤ n + n, by intro h; rw [nat.add_right_comm m m 1, add_assoc]; exact add_le_add h h @@ -184,11 +184,11 @@ namespace pos_num have := cmp_dec a b, revert this, cases cmp a b; dsimp; intro, { exact nat.le_succ_of_le (@add_lt_add nat _ _ _ _ _ this this) }, { rw this, apply nat.lt_succ_self }, - { exact cmp_dec_lemma this } + { exact cmp_dec_theorem this } end | (bit1 a) (bit0 b) := begin dsimp [cmp], have := cmp_dec a b, revert this, cases cmp a b; dsimp; intro, - { exact cmp_dec_lemma this }, + { exact cmp_dec_theorem this }, { rw this, apply nat.lt_succ_self }, { exact nat.le_succ_of_le (@add_lt_add nat _ _ _ _ _ this this) }, end @@ -360,7 +360,7 @@ namespace num decidable_le := num.decidable_le, decidable_eq := num.decidable_eq } - lemma bitwise_to_nat {f : num → num → num} {g : bool → bool → bool} + theorem bitwise_to_nat {f : num → num → num} {g : bool → bool → bool} (p : pos_num → pos_num → num) (gff : g ff ff = ff) (f00 : f 0 0 = 0) @@ -407,23 +407,23 @@ namespace num rw [←bit_to_nat, pbb] } } end - @[simp] lemma lor_to_nat : ∀ m n, (lor m n : ℕ) = nat.lor m n := + @[simp] theorem lor_to_nat : ∀ m n, (lor m n : ℕ) = nat.lor m n := by apply bitwise_to_nat (λx y, ↑(pos_num.lor x y)); intros; try {cases a}; try {cases b}; refl - @[simp] lemma land_to_nat : ∀ m n, (land m n : ℕ) = nat.land m n := + @[simp] theorem land_to_nat : ∀ m n, (land m n : ℕ) = nat.land m n := by apply bitwise_to_nat pos_num.land; intros; try {cases a}; try {cases b}; refl - @[simp] lemma ldiff_to_nat : ∀ m n, (ldiff m n : ℕ) = nat.ldiff m n := + @[simp] theorem ldiff_to_nat : ∀ m n, (ldiff m n : ℕ) = nat.ldiff m n := by apply bitwise_to_nat pos_num.ldiff; intros; try {cases a}; try {cases b}; refl - @[simp] lemma lxor_to_nat : ∀ m n, (lxor m n : ℕ) = nat.lxor m n := + @[simp] theorem lxor_to_nat : ∀ m n, (lxor m n : ℕ) = nat.lxor m n := by apply bitwise_to_nat pos_num.lxor; intros; try {cases a}; try {cases b}; refl - @[simp] lemma shiftl_to_nat (m n) : (shiftl m n : ℕ) = nat.shiftl m n := + @[simp] theorem shiftl_to_nat (m n) : (shiftl m n : ℕ) = nat.shiftl m n := begin cases m; dunfold shiftl, {symmetry, apply nat.zero_shiftl}, induction n with n IH, {refl}, simp [pos_num.shiftl, nat.shiftl_succ], rw ←IH, refl end - @[simp] lemma shiftr_to_nat (m n) : (shiftr m n : ℕ) = nat.shiftr m n := + @[simp] theorem shiftr_to_nat (m n) : (shiftr m n : ℕ) = nat.shiftr m n := begin cases m with m; dunfold shiftr, {symmetry, apply nat.zero_shiftr}, revert m; induction n with n IH; intro m, {cases m; refl}, @@ -444,7 +444,7 @@ namespace num rw nat.div2_bit } end - @[simp] lemma test_bit_to_nat (m n) : test_bit m n = nat.test_bit m n := + @[simp] theorem test_bit_to_nat (m n) : test_bit m n = nat.test_bit m n := begin cases m with m; unfold test_bit nat.test_bit, { change (zero : nat) with 0, rw nat.zero_shiftr, refl }, diff --git a/data/rat.lean b/data/rat.lean index cbd2b9c4e47b1..51eddb1d68cc2 100644 --- a/data/rat.lean +++ b/data/rat.lean @@ -20,13 +20,13 @@ if h₁ : a = b then h_eq h₁ else variables {a b : α} {h_eq : a = b → β} {h_lt : a < b → β} {h_gt : a > b → β} -lemma linear_order_cases_on_eq (h : a = b) : linear_order_cases_on a b h_eq h_lt h_gt = h_eq h := +theorem linear_order_cases_on_eq (h : a = b) : linear_order_cases_on a b h_eq h_lt h_gt = h_eq h := dif_pos h -lemma linear_order_cases_on_lt (h : a < b) : linear_order_cases_on a b h_eq h_lt h_gt = h_lt h := +theorem linear_order_cases_on_lt (h : a < b) : linear_order_cases_on a b h_eq h_lt h_gt = h_lt h := eq.trans (dif_neg $ ne_of_lt h) $ dif_pos h -lemma linear_order_cases_on_gt (h : a > b) : linear_order_cases_on a b h_eq h_lt h_gt = h_gt h := +theorem linear_order_cases_on_gt (h : a > b) : linear_order_cases_on a b h_eq h_lt h_gt = h_gt h := eq.trans (dif_neg $ (ne_of_lt h).symm) (dif_neg $ not_lt_of_ge $ le_of_lt h) end linear_order_cases_on @@ -37,14 +37,14 @@ section ordered_ring universes u variables {α : Type u} [linear_ordered_ring α] {a b : α} -lemma mul_nonneg_iff_right_nonneg_of_pos (h : 0 < a) : 0 ≤ b * a ↔ 0 ≤ b := +theorem mul_nonneg_iff_right_nonneg_of_pos (h : 0 < a) : 0 ≤ b * a ↔ 0 ≤ b := ⟨assume : 0 ≤ b * a, nonneg_of_mul_nonneg_right this h, assume : 0 ≤ b, mul_nonneg this $ le_of_lt h⟩ end ordered_ring /- auxiliary -/ -lemma not_antimono {a b : Prop} (nb : ¬ b) (h : (a → b)) : ¬ a := +theorem not_antimono {a b : Prop} (nb : ¬ b) (h : (a → b)) : ¬ a := assume ha, nb (h ha) /- rational numbers -/ @@ -178,47 +178,47 @@ instance : has_inv ℚ := ⟨rat.inv⟩ variables (a b c : ℚ) -protected lemma add_zero : a + 0 = a := +protected theorem add_zero : a + 0 = a := quotient.induction_on a $ λ⟨n, ⟨d, h⟩⟩, quotient.sound $ by simp [rat.rel] -protected lemma zero_add : 0 + a = a := +protected theorem zero_add : 0 + a = a := quotient.induction_on a $ λ⟨n, ⟨d, h⟩⟩, quotient.sound $ by simp [rat.rel] -protected lemma add_comm : a + b = b + a := +protected theorem add_comm : a + b = b + a := quotient.induction_on₂ a b $ λ⟨n₁, ⟨d₁, h₁⟩⟩ ⟨n₂, ⟨d₂, h₂⟩⟩, quotient.sound $ by simp [rat.rel] -protected lemma add_assoc : a + b + c = a + (b + c) := +protected theorem add_assoc : a + b + c = a + (b + c) := quotient.induction_on₃ a b c $ λ⟨n₁, ⟨d₁, h₁⟩⟩ ⟨n₂, ⟨d₂, h₂⟩⟩ ⟨n₃, ⟨d₃, h₃⟩⟩, quotient.sound $ by simp [rat.rel, mul_add, add_mul] -protected lemma add_left_neg : -a + a = 0 := +protected theorem add_left_neg : -a + a = 0 := quotient.induction_on a $ λ⟨n, ⟨d, h⟩⟩, quotient.sound $ by simp [rat.rel] -protected lemma mul_one : a * 1 = a := +protected theorem mul_one : a * 1 = a := quotient.induction_on a $ λ⟨n, ⟨d, h⟩⟩, quotient.sound $ by simp [rat.rel] -protected lemma one_mul : 1 * a = a := +protected theorem one_mul : 1 * a = a := quotient.induction_on a $ λ⟨n, ⟨d, h⟩⟩, quotient.sound $ by simp [rat.rel] -protected lemma mul_comm : a * b = b * a := +protected theorem mul_comm : a * b = b * a := quotient.induction_on₂ a b $ λ⟨n₁, ⟨d₁, h₁⟩⟩ ⟨n₂, ⟨d₂, h₂⟩⟩, quotient.sound $ by simp [rat.rel] -protected lemma mul_assoc : a * b * c = a * (b * c) := +protected theorem mul_assoc : a * b * c = a * (b * c) := quotient.induction_on₃ a b c $ λ⟨n₁, ⟨d₁, h₁⟩⟩ ⟨n₂, ⟨d₂, h₂⟩⟩ ⟨n₃, ⟨d₃, h₃⟩⟩, quotient.sound $ by simp [rat.rel] -protected lemma add_mul : (a + b) * c = a * c + b * c := +protected theorem add_mul : (a + b) * c = a * c + b * c := quotient.induction_on₃ a b c $ λ⟨n₁, ⟨d₁, h₁⟩⟩ ⟨n₂, ⟨d₂, h₂⟩⟩ ⟨n₃, ⟨d₃, h₃⟩⟩, quotient.sound $ by simp [rat.rel, mul_add, add_mul] -protected lemma mul_add : a * (b + c) = a * b + a * c := +protected theorem mul_add : a * (b + c) = a * b + a * c := quotient.induction_on₃ a b c $ λ⟨n₁, ⟨d₁, h₁⟩⟩ ⟨n₂, ⟨d₂, h₂⟩⟩ ⟨n₃, ⟨d₃, h₃⟩⟩, quotient.sound $ by simp [rat.rel, mul_add, add_mul] @@ -233,10 +233,10 @@ private lemma eq_zero_of_rat_eq_zero : ∀{a : rat.num_denum}, a.1 = 0 → ⟦a private lemma rat_eq_zero_iff {a : rat.num_denum} : ⟦a⟧ = (0:ℚ) ↔ a.1 = 0 := ⟨rat_eq_zero, eq_zero_of_rat_eq_zero⟩ -protected lemma zero_ne_one : 0 ≠ (1:ℚ) := +protected theorem zero_ne_one : 0 ≠ (1:ℚ) := assume h, zero_ne_one (rat_eq_zero h.symm).symm -protected lemma mul_inv_cancel : a ≠ 0 → a * a⁻¹ = 1 := +protected theorem mul_inv_cancel : a ≠ 0 → a * a⁻¹ = 1 := quotient.induction_on a $ λ⟨n, ⟨d, h⟩⟩ neq0, let a : rat.num_denum := ⟨n, ⟨d, h⟩⟩ in linear_order_cases_on n 0 @@ -250,7 +250,7 @@ linear_order_cases_on n 0 from @inv'_pos n d h _, begin simp [this], apply quotient.sound, simp [rat.rel] end) -protected lemma inv_mul_cancel (h : a ≠ 0) : a⁻¹ * a = 1 := +protected theorem inv_mul_cancel (h : a ≠ 0) : a⁻¹ * a = 1 := eq.trans (rat.mul_comm _ _) (rat.mul_inv_cancel _ h) instance decidable_eq_rat.rel : Π{a b : rat.num_denum}, decidable (rat.rel a b) @@ -318,7 +318,7 @@ protected def le (a b : ℚ) := rat.nonneg (b - a) instance : has_le ℚ := ⟨rat.le⟩ -protected lemma le_refl : a ≤ a := +protected theorem le_refl : a ≤ a := show rat.nonneg (a - a), begin rw [sub_self], exact le_refl (0 : int) end instance : linear_strong_order_pair ℚ := diff --git a/data/seq/computation.lean b/data/seq/computation.lean index 8f9f739fc8148..21f6f473d5fdd 100644 --- a/data/seq/computation.lean +++ b/data/seq/computation.lean @@ -170,7 +170,7 @@ section bisim def is_bisimulation := ∀ ⦃s₁ s₂⦄, s₁ ~ s₂ → bisim_o R (destruct s₁) (destruct s₂) -- If two computations are bisimilar, then they are equal - lemma eq_of_bisim (bisim : is_bisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s₁ = s₂ := + theorem eq_of_bisim (bisim : is_bisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s₁ = s₂ := begin apply subtype.eq, apply stream.eq_of_bisim (λx y, ∃ s s' : computation α, s.1 = x ∧ s'.1 = y ∧ R s s'), @@ -456,12 +456,12 @@ theorem has_bind_eq_bind {β} (c : computation α) (f : α → computation β) : def join (c : computation (computation α)) : computation α := c >>= id -@[simp] lemma map_ret (f : α → β) (a) : map f (return a) = return (f a) := rfl +@[simp] theorem map_ret (f : α → β) (a) : map f (return a) = return (f a) := rfl -@[simp] lemma map_think (f : α → β) : ∀ s, map f (think s) = think (map f s) +@[simp] theorem map_think (f : α → β) : ∀ s, map f (think s) = think (map f s) | ⟨s, al⟩ := by apply subtype.eq; dsimp [think, map]; rw stream.map_cons -@[simp] lemma destruct_map (f : α → β) (s) : destruct (map f s) = lmap f (rmap (map f) (destruct s)) := +@[simp] theorem destruct_map (f : α → β) (s) : destruct (map f s) = lmap f (rmap (map f) (destruct s)) := by apply s.cases_on; intro; simp @[simp] theorem map_id : ∀ (s : computation α), map id s = s @@ -472,7 +472,7 @@ by apply s.cases_on; intro; simp simp [e, stream.map_id] end -lemma map_comp (f : α → β) (g : β → γ) : +theorem map_comp (f : α → β) (g : β → γ) : ∀ (s : computation α), map (g ∘ f) s = map g (map f s) | ⟨s, al⟩ := begin apply subtype.eq; dsimp [map], @@ -481,7 +481,7 @@ lemma map_comp (f : α → β) (g : β → γ) : apply funext, intro, cases x with x; refl end -@[simp] lemma ret_bind (a) (f : α → computation β) : +@[simp] theorem ret_bind (a) (f : α → computation β) : bind (return a) f = f a := begin apply eq_of_bisim (λc1 c2, @@ -500,7 +500,7 @@ begin { simp } end -@[simp] lemma think_bind (c) (f : α → computation β) : +@[simp] theorem think_bind (c) (f : α → computation β) : bind (think c) f = think (bind c f) := destruct_eq_think $ by simp [bind, bind.F] @@ -606,10 +606,10 @@ instance : monad computation := theorem has_map_eq_map {β} (f : α → β) (c : computation α) : f <$> c = map f c := rfl -@[simp] lemma return_def (a) : (_root_.return a : computation α) = return a := rfl +@[simp] theorem return_def (a) : (_root_.return a : computation α) = return a := rfl -@[simp] lemma map_ret' {α β} : ∀ (f : α → β) (a), f <$> return a = return (f a) := map_ret -@[simp] lemma map_think' {α β} : ∀ (f : α → β) s, f <$> think s = think (f <$> s) := map_think +@[simp] theorem map_ret' {α β} : ∀ (f : α → β) (a), f <$> return a = return (f a) := map_ret +@[simp] theorem map_think' {α β} : ∀ (f : α → β) s, f <$> think s = think (f <$> s) := map_think theorem mem_map (f : α → β) {a} {s : computation α} (m : a ∈ s) : f a ∈ map f s := by rw ←bind_ret; apply mem_bind m; apply ret_mem @@ -884,7 +884,7 @@ by cases a with a ca; cases b with b cb; simp only [lift_rel_aux] lift_rel_aux R C (destruct ca) (sum.inl b) ↔ ∃ {a}, a ∈ ca ∧ R a b := by rw [←lift_rel_aux.swap, lift_rel_aux.ret_left] -lemma lift_rel_rec.lem {R : α → β → Prop} (C : computation α → computation β → Prop) +theorem lift_rel_rec.lem {R : α → β → Prop} (C : computation α → computation β → Prop) (H : ∀ {ca cb}, C ca cb → lift_rel_aux R C (destruct ca) (destruct cb)) (ca cb) (Hc : C ca cb) (a) (ha : a ∈ ca) : lift_rel R ca cb := begin diff --git a/data/seq/parallel.lean b/data/seq/parallel.lean index 69b2413719294..4e2a56f85e98c 100644 --- a/data/seq/parallel.lean +++ b/data/seq/parallel.lean @@ -36,7 +36,7 @@ def parallel.aux1 : list (computation α) × wseq (computation α) → def parallel (S : wseq (computation α)) : computation α := corec parallel.aux1 ([], S) -lemma terminates_parallel.aux : ∀ {l : list (computation α)} {S c}, +theorem terminates_parallel.aux : ∀ {l : list (computation α)} {S c}, c ∈ l → terminates c → terminates (corec parallel.aux1 (l, S)) := begin have lem1 : ∀ l S, (∃ (a : α), parallel.aux2 l = sum.inl a) → @@ -176,7 +176,7 @@ begin exact seq.mem_cons_of_mem _ dS' } } } end -lemma map_parallel (f : α → β) (S) : map f (parallel S) = parallel (S.map (map f)) := +theorem map_parallel (f : α → β) (S) : map f (parallel S) = parallel (S.map (map f)) := begin refine eq_of_bisim (λ c1 c2, ∃ l S, c1 = map f (corec parallel.aux1 (l, S)) ∧ @@ -233,7 +233,7 @@ theorem mem_parallel {S : wseq (computation α)} {a} by have := terminates_of_mem ac; have := terminates_parallel cs; exact mem_of_promises _ (parallel_promises H) -lemma parallel_congr_lem {S T : wseq (computation α)} {a} +theorem parallel_congr_lem {S T : wseq (computation α)} {a} (H : S.lift_rel equiv T) : (∀ s ∈ S, s ~> a) ↔ (∀ t ∈ T, t ~> a) := ⟨λ h1 t tT, let ⟨s, sS, se⟩ := wseq.exists_of_lift_rel_right H tT in (promises_congr se _).1 (h1 _ sS), diff --git a/data/seq/seq.lean b/data/seq/seq.lean index 067bb2e8263a4..7c0341bbfa4d4 100644 --- a/data/seq/seq.lean +++ b/data/seq/seq.lean @@ -43,17 +43,17 @@ by {cases s with f al, induction h with n h IH, exacts [id, λ h2, al (IH h2)]} theorem not_mem_nil (a : α) : a ∉ @nil α := λ ⟨n, (h : some a = none)⟩, by injection h -lemma mem_cons (a : α) : ∀ (s : seq α), a ∈ cons a s +theorem mem_cons (a : α) : ∀ (s : seq α), a ∈ cons a s | ⟨f, al⟩ := stream.mem_cons (some a) _ -lemma mem_cons_of_mem (y : α) {a : α} : ∀ {s : seq α}, a ∈ s → a ∈ cons y s +theorem mem_cons_of_mem (y : α) {a : α} : ∀ {s : seq α}, a ∈ s → a ∈ cons y s | ⟨f, al⟩ := stream.mem_cons_of_mem (some y) -lemma eq_or_mem_of_mem_cons {a b : α} : ∀ {s : seq α}, a ∈ cons b s → a = b ∨ a ∈ s +theorem eq_or_mem_of_mem_cons {a b : α} : ∀ {s : seq α}, a ∈ cons b s → a = b ∨ a ∈ s | ⟨f, al⟩ h := or_of_or_of_implies_left (stream.eq_or_mem_of_mem_cons h) (λh, by injection h) -@[simp] lemma mem_cons_iff {a b : α} {s : seq α} : a ∈ cons b s ↔ a = b ∨ a ∈ s := +@[simp] theorem mem_cons_iff {a b : α} {s : seq α} : a ∈ cons b s ↔ a = b ∨ a ∈ s := ⟨eq_or_mem_of_mem_cons, λo, by cases o with e m; [{rw e, apply mem_cons}, exact mem_cons_of_mem _ m]⟩ @@ -187,7 +187,7 @@ section bisim def is_bisimulation := ∀ ⦃s₁ s₂⦄, s₁ ~ s₂ → bisim_o R (destruct s₁) (destruct s₂) -- If two streams are bisimilar, then they are equal - lemma eq_of_bisim (bisim : is_bisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s₁ = s₂ := + theorem eq_of_bisim (bisim : is_bisimulation R) {s₁ s₂} (r : s₁ ~ s₂) : s₁ = s₂ := begin apply subtype.eq, apply stream.eq_of_bisim (λx y, ∃ s s' : seq α, s.1 = x ∧ s'.1 = y ∧ R s s'), @@ -215,13 +215,13 @@ section bisim end end bisim -lemma coinduction : ∀ {s₁ s₂ : seq α}, head s₁ = head s₂ → +theorem coinduction : ∀ {s₁ s₂ : seq α}, head s₁ = head s₂ → (∀ (β : Type u) (fr : seq α → β), fr s₁ = fr s₂ → fr (tail s₁) = fr (tail s₂)) → s₁ = s₂ | ⟨f₁, a₁⟩ ⟨f₂, a₂⟩ hh ht := subtype.eq (stream.coinduction hh (λ β fr, ht β (λs, fr s.1))) -lemma coinduction2 (s) (f g : seq α → seq β) +theorem coinduction2 (s) (f g : seq α → seq β) (H : ∀ s, bisim_o (λ (s1 s2 : seq β), ∃ (s : seq α), s1 = f s ∧ s2 = g s) (destruct (f s)) (destruct (g s))) : f s = g s := @@ -325,7 +325,7 @@ if h : ∃ n, ¬ (nth s n).is_some then sum.inl (to_list s h) else sum.inr (to_stream s (λn, decidable.by_contradiction (λ hn, h ⟨n, hn⟩))) -@[simp] lemma nil_append (s : seq α) : append nil s = s := +@[simp] theorem nil_append (s : seq α) : append nil s = s := begin apply coinduction2, intro s, dsimp [append], rw [corec_eq], @@ -336,14 +336,14 @@ begin exact ⟨rfl, s, rfl, rfl⟩ } end -@[simp] lemma cons_append (a : α) (s t) : append (cons a s) t = cons a (append s t) := +@[simp] theorem cons_append (a : α) (s t) : append (cons a s) t = cons a (append s t) := destruct_eq_cons $ begin dsimp [append], rw [corec_eq], dsimp [append], rw [destruct_cons], dsimp [append], refl end -@[simp] lemma append_nil (s : seq α) : append s nil = s := +@[simp] theorem append_nil (s : seq α) : append s nil = s := begin apply coinduction2 s, intro s, apply cases_on s _ _, @@ -353,7 +353,7 @@ begin exact ⟨rfl, s, rfl, rfl⟩ } end -@[simp] lemma append_assoc (s t u : seq α) : +@[simp] theorem append_assoc (s t u : seq α) : append (append s t) u = append s (append t u) := begin apply eq_of_bisim (λs1 s2, ∃ s t u, @@ -369,21 +369,21 @@ begin { exact ⟨s, t, u, rfl, rfl⟩ } end -@[simp] lemma map_nil (f : α → β) : map f nil = nil := rfl +@[simp] theorem map_nil (f : α → β) : map f nil = nil := rfl -@[simp] lemma map_cons (f : α → β) (a) : ∀ s, map f (cons a s) = cons (f a) (map f s) +@[simp] theorem map_cons (f : α → β) (a) : ∀ s, map f (cons a s) = cons (f a) (map f s) | ⟨s, al⟩ := by apply subtype.eq; dsimp [cons, map]; rw stream.map_cons; refl -@[simp] lemma map_id : ∀ (s : seq α), map id s = s +@[simp] theorem map_id : ∀ (s : seq α), map id s = s | ⟨s, al⟩ := begin apply subtype.eq; dsimp [map], rw [option.map_id, stream.map_id]; refl end -@[simp] lemma map_tail (f : α → β) : ∀ s, map f (tail s) = tail (map f s) +@[simp] theorem map_tail (f : α → β) : ∀ s, map f (tail s) = tail (map f s) | ⟨s, al⟩ := by apply subtype.eq; dsimp [tail, map]; rw stream.map_tail; refl -lemma map_comp (f : α → β) (g : β → γ) : ∀ (s : seq α), map (g ∘ f) s = map g (map f s) +theorem map_comp (f : α → β) (g : β → γ) : ∀ (s : seq α), map (g ∘ f) s = map g (map f s) | ⟨s, al⟩ := begin apply subtype.eq; dsimp [map], rw stream.map_map, @@ -403,23 +403,23 @@ begin end end end -@[simp] lemma map_nth (f : α → β) : ∀ s n, nth (map f s) n = option_map f (nth s n) +@[simp] theorem map_nth (f : α → β) : ∀ s n, nth (map f s) n = option_map f (nth s n) | ⟨s, al⟩ n := rfl instance : functor seq := { map := @map, id_map := @map_id, map_comp := @map_comp } -@[simp] lemma join_nil : join nil = (nil : seq α) := destruct_eq_nil rfl +@[simp] theorem join_nil : join nil = (nil : seq α) := destruct_eq_nil rfl -@[simp] lemma join_cons_nil (a : α) (S) : +@[simp] theorem join_cons_nil (a : α) (S) : join (cons (a, nil) S) = cons a (join S) := destruct_eq_cons $ by simp [join] -@[simp] lemma join_cons_cons (a b : α) (s S) : +@[simp] theorem join_cons_cons (a b : α) (s S) : join (cons (a, cons b s) S) = cons a (join (cons (b, s) S)) := destruct_eq_cons $ by simp [join] -@[simp] lemma join_cons (a : α) (s S) : +@[simp] theorem join_cons (a : α) (s S) : join (cons (a, s) S) = cons a (append s (join S)) := begin apply eq_of_bisim (λs1 s2, s1 = s2 ∨ @@ -439,7 +439,7 @@ begin end end -@[simp] lemma join_append (S T : seq (seq1 α)) : +@[simp] theorem join_append (S T : seq (seq1 α)) : join (append S T) = append (join S) (join T) := begin apply eq_of_bisim (λs1 s2, ∃ s S T, @@ -505,7 +505,7 @@ end theorem mem_map (f : α → β) {a : α} : ∀ {s : seq α}, a ∈ s → f a ∈ map f s | ⟨g, al⟩ := stream.mem_map (option_map f) -lemma exists_of_mem_map {f} {b : β} : ∀ {s : seq α}, b ∈ map f s → ∃ a, a ∈ s ∧ f a = b +theorem exists_of_mem_map {f} {b : β} : ∀ {s : seq α}, b ∈ map f s → ∃ a, a ∈ s ∧ f a = b | ⟨g, al⟩ h := let ⟨o, om, oe⟩ := stream.exists_of_mem_map h in by cases o; injection oe; exact ⟨a, om, h⟩ @@ -550,9 +550,9 @@ def join : seq1 (seq1 α) → seq1 α | some s' := (a, seq.join (cons s' S)) end -@[simp] lemma join_nil (a : α) (S) : join ((a, nil), S) = (a, seq.join S) := rfl +@[simp] theorem join_nil (a : α) (S) : join ((a, nil), S) = (a, seq.join S) := rfl -@[simp] lemma join_cons (a b : α) (s S) : +@[simp] theorem join_cons (a b : α) (s S) : join ((a, cons b s), S) = (a, seq.join (cons (b, s) S)) := by dsimp [join]; rw [destruct_cons]; refl diff --git a/data/seq/wseq.lean b/data/seq/wseq.lean index 5e52d359347ed..8b0799f164b35 100644 --- a/data/seq/wseq.lean +++ b/data/seq/wseq.lean @@ -475,17 +475,17 @@ by { simp [think, join], unfold has_map.map, simp [join, seq1.ret] } join (cons s S) = think (append s (join S)) := by { simp [think, join], unfold has_map.map, simp [join, cons, append] } -@[simp] lemma nil_append (s : wseq α) : append nil s = s := seq.nil_append _ +@[simp] theorem nil_append (s : wseq α) : append nil s = s := seq.nil_append _ -@[simp] lemma cons_append (a : α) (s t) : +@[simp] theorem cons_append (a : α) (s t) : append (cons a s) t = cons a (append s t) := seq.cons_append _ _ _ -@[simp] lemma think_append (s t : wseq α) : +@[simp] theorem think_append (s t : wseq α) : append (think s) t = think (append s t) := seq.cons_append _ _ _ -@[simp] lemma append_nil (s : wseq α) : append s nil = s := seq.append_nil _ +@[simp] theorem append_nil (s : wseq α) : append s nil = s := seq.append_nil _ -@[simp] lemma append_assoc (s t u : wseq α) : +@[simp] theorem append_assoc (s t u : wseq α) : append (append s t) u = append s (append t u) := seq.append_assoc _ _ _ def tail.aux : option (α × wseq α) → computation (option (α × wseq α)) @@ -710,7 +710,7 @@ seq.of_mem_append def mem_append_left {s₁ s₂ : wseq α} {a : α} : a ∈ s₁ → a ∈ append s₁ s₂ := seq.mem_append_left -lemma exists_of_mem_map {f} {b : β} : ∀ {s : wseq α}, b ∈ map f s → ∃ a, a ∈ s ∧ f a = b +theorem exists_of_mem_map {f} {b : β} : ∀ {s : wseq α}, b ∈ map f s → ∃ a, a ∈ s ∧ f a = b | ⟨g, al⟩ h := let ⟨o, om, oe⟩ := seq.exists_of_mem_map h in by cases o; injection oe; exact ⟨a, om, h⟩ @@ -922,12 +922,12 @@ end def ret (a : α) : wseq α := of_list [a] -@[simp] lemma map_nil (f : α → β) : map f nil = nil := rfl +@[simp] theorem map_nil (f : α → β) : map f nil = nil := rfl -@[simp] lemma map_cons (f : α → β) (a s) : +@[simp] theorem map_cons (f : α → β) (a s) : map f (cons a s) = cons (f a) (map f s) := seq.map_cons _ _ _ -@[simp] lemma map_think (f : α → β) (s) : +@[simp] theorem map_think (f : α → β) (s) : map f (think s) = think (map f s) := seq.map_cons _ _ _ @[simp] theorem map_id (s : wseq α) : map id s = s := by simp [map] @@ -937,7 +937,7 @@ def ret (a : α) : wseq α := of_list [a] @[simp] theorem map_append (f : α → β) (s t) : map f (append s t) = append (map f s) (map f t) := seq.map_append _ _ _ -lemma map_comp (f : α → β) (g : β → γ) (s : wseq α) : +theorem map_comp (f : α → β) (g : β → γ) (s : wseq α) : map (g ∘ f) s = map g (map f s) := begin dsimp [map], rw ←seq.map_comp, @@ -982,7 +982,7 @@ theorem exists_of_mem_bind {s : wseq α} {f : α → wseq β} {b} let ⟨t, tm, bt⟩ := exists_of_mem_join h, ⟨a, as, e⟩ := exists_of_mem_map tm in ⟨a, as, by rwa e⟩ -lemma destruct_map (f : α → β) (s : wseq α) : +theorem destruct_map (f : α → β) (s : wseq α) : destruct (map f s) = computation.map (option_map (prod.map f (map f))) (destruct s) := begin apply eq_of_bisim (λ c1 c2, ∃ s, c1 = destruct (map f s) ∧ @@ -1019,7 +1019,7 @@ def destruct_append.aux (t : wseq α) : | (some (a, s)) := return (some (a, append s t)) attribute [simp] destruct_append.aux -lemma destruct_append (s t : wseq α) : +theorem destruct_append (s t : wseq α) : destruct (append s t) = (destruct s).bind (destruct_append.aux t) := begin apply eq_of_bisim (λ c1 c2, ∃ s t, c1 = destruct (append s t) ∧ @@ -1036,7 +1036,7 @@ def destruct_join.aux : option (wseq α × wseq (wseq α)) → computation (opti | (some (s, S)) := (destruct (append s (join S))).think attribute [simp] destruct_join.aux -lemma destruct_join (S : wseq (wseq α)) : +theorem destruct_join (S : wseq (wseq α)) : destruct (join S) = (destruct S).bind destruct_join.aux := begin apply eq_of_bisim (λ c1 c2, c1 = c2 ∨ ∃ S, c1 = destruct (join S) ∧ @@ -1175,7 +1175,7 @@ begin { exact ⟨s, rfl, rfl⟩ } end -@[simp] lemma join_append (S T : wseq (wseq α)) : +@[simp] theorem join_append (S T : wseq (wseq α)) : join (append S T) ~ append (join S) (join T) := begin refine ⟨λ s1 s2, ∃ s S T, diff --git a/data/set/basic.lean b/data/set/basic.lean index 6f06499545701..5a415c1a52a9d 100644 --- a/data/set/basic.lean +++ b/data/set/basic.lean @@ -62,7 +62,7 @@ theorem empty_def : (∅ : set α) = {x | false} := rfl @[simp] theorem mem_empty_eq (x : α) : x ∈ (∅ : set α) = false := rfl -@[simp] lemma set_of_false : {a : α | false} = ∅ := rfl +@[simp] theorem set_of_false : {a : α | false} = ∅ := rfl theorem eq_empty_of_forall_not_mem {s : set α} (h : ∀ x, x ∉ s) : s = ∅ := ext (assume x, iff.intro diff --git a/data/set/finite.lean b/data/set/finite.lean index 91c872422b624..9a4daabcedde9 100644 --- a/data/set/finite.lean +++ b/data/set/finite.lean @@ -27,19 +27,19 @@ inductive finite : set α → Prop attribute [simp] finite.empty @[simp] -lemma finite_insert {a : α} {s : set α} (h : finite s) : finite (insert a s) := +theorem finite_insert {a : α} {s : set α} (h : finite s) : finite (insert a s) := classical.by_cases (assume : a ∈ s, by simp [*]) (assume : a ∉ s, finite.insert a s this h) @[simp] -lemma finite_singleton {a : α} : finite ({a} : set α) := +theorem finite_singleton {a : α} : finite ({a} : set α) := finite_insert finite.empty -lemma finite_union {s t : set α} (hs : finite s) (ht : finite t) : finite (s ∪ t) := +theorem finite_union {s t : set α} (hs : finite s) (ht : finite t) : finite (s ∪ t) := finite.drec_on ht (by simp [hs]) $ assume a t _ _, by simp; exact finite_insert -lemma finite_subset {s : set α} (hs : finite s) : ∀{t}, t ⊆ s → finite t := +theorem finite_subset {s : set α} (hs : finite s) : ∀{t}, t ⊆ s → finite t := begin induction hs with a t' ha ht' ih, { intros t ht, simp [(subset_empty_iff t).mp ht, finite.empty] }, @@ -54,14 +54,14 @@ begin { simp [sdiff_singleton_eq_same, hna] at tm, exact tm } } end -lemma finite_image {s : set α} {f : α → β} (h : finite s) : finite (f '' s) := +theorem finite_image {s : set α} {f : α → β} (h : finite s) : finite (f '' s) := begin induction h with a s' hns' hs' hi, simp [image_empty, finite.empty], simp [image_insert_eq, finite_insert, hi] end -lemma finite_sUnion {s : set (set α)} (h : finite s) : (∀t∈s, finite t) → finite (⋃₀ s) := +theorem finite_sUnion {s : set (set α)} (h : finite s) : (∀t∈s, finite t) → finite (⋃₀ s) := begin induction h with a s' hns' hs' hi, { simp [finite.empty] }, diff --git a/data/stream.lean b/data/stream.lean index 4d886615f0b2a..b92efab8767f3 100644 --- a/data/stream.lean +++ b/data/stream.lean @@ -32,47 +32,47 @@ def drop (n : nat) (s : stream α) : stream α := @[reducible] def nth (n : nat) (s : stream α) : α := s n -protected lemma eta (s : stream α) : head s :: tail s = s := +protected theorem eta (s : stream α) : head s :: tail s = s := funext (λ i, begin cases i, repeat {reflexivity} end) -lemma nth_zero_cons (a : α) (s : stream α) : nth 0 (a :: s) = a := +theorem nth_zero_cons (a : α) (s : stream α) : nth 0 (a :: s) = a := rfl -lemma head_cons (a : α) (s : stream α) : head (a :: s) = a := +theorem head_cons (a : α) (s : stream α) : head (a :: s) = a := rfl -lemma tail_cons (a : α) (s : stream α) : tail (a :: s) = s := +theorem tail_cons (a : α) (s : stream α) : tail (a :: s) = s := rfl -lemma tail_drop (n : nat) (s : stream α) : tail (drop n s) = drop n (tail s) := +theorem tail_drop (n : nat) (s : stream α) : tail (drop n s) = drop n (tail s) := funext (λ i, begin unfold tail drop, simp end) -lemma nth_drop (n m : nat) (s : stream α) : nth n (drop m s) = nth (n+m) s := +theorem nth_drop (n m : nat) (s : stream α) : nth n (drop m s) = nth (n+m) s := rfl -lemma tail_eq_drop (s : stream α) : tail s = drop 1 s := +theorem tail_eq_drop (s : stream α) : tail s = drop 1 s := rfl -lemma drop_drop (n m : nat) (s : stream α) : drop n (drop m s) = drop (n+m) s := +theorem drop_drop (n m : nat) (s : stream α) : drop n (drop m s) = drop (n+m) s := funext (λ i, begin unfold drop, rw add_assoc end) -lemma nth_succ (n : nat) (s : stream α) : nth (succ n) s = nth n (tail s) := +theorem nth_succ (n : nat) (s : stream α) : nth (succ n) s = nth n (tail s) := rfl -lemma drop_succ (n : nat) (s : stream α) : drop (succ n) s = drop n (tail s) := +theorem drop_succ (n : nat) (s : stream α) : drop (succ n) s = drop n (tail s) := rfl -protected lemma ext {s₁ s₂ : stream α} : (∀ n, nth n s₁ = nth n s₂) → s₁ = s₂ := +protected theorem ext {s₁ s₂ : stream α} : (∀ n, nth n s₁ = nth n s₂) → s₁ = s₂ := assume h, funext h def all (p : α → Prop) (s : stream α) := ∀ n, p (nth n s) def any (p : α → Prop) (s : stream α) := ∃ n, p (nth n s) -lemma all_def (p : α → Prop) (s : stream α) : all p s = ∀ n, p (nth n s) := +theorem all_def (p : α → Prop) (s : stream α) : all p s = ∀ n, p (nth n s) := rfl -lemma any_def (p : α → Prop) (s : stream α) : any p s = ∃ n, p (nth n s) := +theorem any_def (p : α → Prop) (s : stream α) : any p s = ∃ n, p (nth n s) := rfl protected def mem (a : α) (s : stream α) := any (λ b, a = b) s @@ -80,14 +80,14 @@ protected def mem (a : α) (s : stream α) := any (λ b, a = b) s instance : has_mem α (stream α) := ⟨stream.mem⟩ -lemma mem_cons (a : α) (s : stream α) : a ∈ (a::s) := +theorem mem_cons (a : α) (s : stream α) : a ∈ (a::s) := exists.intro 0 rfl -lemma mem_cons_of_mem {a : α} {s : stream α} (b : α) : a ∈ s → a ∈ b :: s := +theorem mem_cons_of_mem {a : α} {s : stream α} (b : α) : a ∈ s → a ∈ b :: s := assume ⟨n, h⟩, exists.intro (succ n) (by rw [nth_succ, tail_cons, h]) -lemma eq_or_mem_of_mem_cons {a b : α} {s : stream α} : a ∈ b::s → a = b ∨ a ∈ s := +theorem eq_or_mem_of_mem_cons {a b : α} {s : stream α} : a ∈ b::s → a = b ∨ a ∈ s := assume ⟨n, h⟩, begin cases n with n', @@ -95,7 +95,7 @@ begin {right, rw [nth_succ, tail_cons] at h, exact ⟨n', h⟩} end -lemma mem_of_nth_eq {n : nat} {s : stream α} {a : α} : a = nth n s → a ∈ s := +theorem mem_of_nth_eq {n : nat} {s : stream α} {a : α} : a = nth n s → a ∈ s := assume h, exists.intro n h section map @@ -104,38 +104,38 @@ variable (f : α → β) def map (s : stream α) : stream β := λ n, f (nth n s) -lemma drop_map (n : nat) (s : stream α) : drop n (map f s) = map f (drop n s) := +theorem drop_map (n : nat) (s : stream α) : drop n (map f s) = map f (drop n s) := stream.ext (λ i, rfl) -lemma nth_map (n : nat) (s : stream α) : nth n (map f s) = f (nth n s) := +theorem nth_map (n : nat) (s : stream α) : nth n (map f s) = f (nth n s) := rfl -lemma tail_map (s : stream α) : tail (map f s) = map f (tail s) := +theorem tail_map (s : stream α) : tail (map f s) = map f (tail s) := begin rw tail_eq_drop, reflexivity end -lemma head_map (s : stream α) : head (map f s) = f (head s) := +theorem head_map (s : stream α) : head (map f s) = f (head s) := rfl -lemma map_eq (s : stream α) : map f s = f (head s) :: map f (tail s) := +theorem map_eq (s : stream α) : map f s = f (head s) :: map f (tail s) := by rw [← stream.eta (map f s), tail_map, head_map] -lemma map_cons (a : α) (s : stream α) : map f (a :: s) = f a :: map f s := +theorem map_cons (a : α) (s : stream α) : map f (a :: s) = f a :: map f s := begin rw [← stream.eta (map f (a :: s)), map_eq], reflexivity end -lemma map_id (s : stream α) : map id s = s := +theorem map_id (s : stream α) : map id s = s := rfl -lemma map_map (g : β → δ) (f : α → β) (s : stream α) : map g (map f s) = map (g ∘ f) s := +theorem map_map (g : β → δ) (f : α → β) (s : stream α) : map g (map f s) = map (g ∘ f) s := rfl -lemma map_tail (s : stream α) : map f (tail s) = tail (map f s) := +theorem map_tail (s : stream α) : map f (tail s) = tail (map f s) := rfl -lemma mem_map {a : α} {s : stream α} : a ∈ s → f a ∈ map f s := +theorem mem_map {a : α} {s : stream α} : a ∈ s → f a ∈ map f s := assume ⟨n, h⟩, exists.intro n (by rw [nth_map, h]) -lemma exists_of_mem_map {f} {b : β} {s : stream α} : b ∈ map f s → ∃ a, a ∈ s ∧ f a = b := +theorem exists_of_mem_map {f} {b : β} {s : stream α} : b ∈ map f s → ∃ a, a ∈ s ∧ f a = b := assume ⟨n, h⟩, ⟨nth n s, ⟨n, rfl⟩, h.symm⟩ end map @@ -145,19 +145,19 @@ variable (f : α → β → δ) def zip (s₁ : stream α) (s₂ : stream β) : stream δ := λ n, f (nth n s₁) (nth n s₂) -lemma drop_zip (n : nat) (s₁ : stream α) (s₂ : stream β) : drop n (zip f s₁ s₂) = zip f (drop n s₁) (drop n s₂) := +theorem drop_zip (n : nat) (s₁ : stream α) (s₂ : stream β) : drop n (zip f s₁ s₂) = zip f (drop n s₁) (drop n s₂) := stream.ext (λ i, rfl) -lemma nth_zip (n : nat) (s₁ : stream α) (s₂ : stream β) : nth n (zip f s₁ s₂) = f (nth n s₁) (nth n s₂) := +theorem nth_zip (n : nat) (s₁ : stream α) (s₂ : stream β) : nth n (zip f s₁ s₂) = f (nth n s₁) (nth n s₂) := rfl -lemma head_zip (s₁ : stream α) (s₂ : stream β) : head (zip f s₁ s₂) = f (head s₁) (head s₂) := +theorem head_zip (s₁ : stream α) (s₂ : stream β) : head (zip f s₁ s₂) = f (head s₁) (head s₂) := rfl -lemma tail_zip (s₁ : stream α) (s₂ : stream β) : tail (zip f s₁ s₂) = zip f (tail s₁) (tail s₂) := +theorem tail_zip (s₁ : stream α) (s₂ : stream β) : tail (zip f s₁ s₂) = zip f (tail s₁) (tail s₂) := rfl -lemma zip_eq (s₁ : stream α) (s₂ : stream β) : zip f s₁ s₂ = f (head s₁) (head s₂) :: zip f (tail s₁) (tail s₂) := +theorem zip_eq (s₁ : stream α) (s₂ : stream β) : zip f s₁ s₂ = f (head s₁) (head s₂) :: zip f (tail s₁) (tail s₂) := begin rw [← stream.eta (zip f s₁ s₂)], reflexivity end end zip @@ -165,35 +165,35 @@ end zip def const (a : α) : stream α := λ n, a -lemma mem_const (a : α) : a ∈ const a := +theorem mem_const (a : α) : a ∈ const a := exists.intro 0 rfl -lemma const_eq (a : α) : const a = a :: const a := +theorem const_eq (a : α) : const a = a :: const a := begin apply stream.ext, intro n, cases n, repeat {reflexivity} end -lemma tail_const (a : α) : tail (const a) = const a := +theorem tail_const (a : α) : tail (const a) = const a := suffices tail (a :: const a) = const a, by rwa [← const_eq] at this, rfl -lemma map_const (f : α → β) (a : α) : map f (const a) = const (f a) := +theorem map_const (f : α → β) (a : α) : map f (const a) = const (f a) := rfl -lemma nth_const (n : nat) (a : α) : nth n (const a) = a := +theorem nth_const (n : nat) (a : α) : nth n (const a) = a := rfl -lemma drop_const (n : nat) (a : α) : drop n (const a) = const a := +theorem drop_const (n : nat) (a : α) : drop n (const a) = const a := stream.ext (λ i, rfl) def iterate (f : α → α) (a : α) : stream α := λ n, nat.rec_on n a (λ n r, f r) -lemma head_iterate (f : α → α) (a : α) : head (iterate f a) = a := +theorem head_iterate (f : α → α) (a : α) : head (iterate f a) = a := rfl -lemma tail_iterate (f : α → α) (a : α) : tail (iterate f a) = iterate f (f a) := +theorem tail_iterate (f : α → α) (a : α) : tail (iterate f a) = iterate f (f a) := begin apply funext, intro n, induction n with n' ih, @@ -204,16 +204,16 @@ begin rw add_one_eq_succ, dsimp, rw ih} end -lemma iterate_eq (f : α → α) (a : α) : iterate f a = a :: iterate f (f a) := +theorem iterate_eq (f : α → α) (a : α) : iterate f a = a :: iterate f (f a) := begin rw [← stream.eta (iterate f a)], rw tail_iterate, reflexivity end -lemma nth_zero_iterate (f : α → α) (a : α) : nth 0 (iterate f a) = a := +theorem nth_zero_iterate (f : α → α) (a : α) : nth 0 (iterate f a) = a := rfl -lemma nth_succ_iterate (n : nat) (f : α → α) (a : α) : nth (succ n) (iterate f a) = nth n (iterate f (f a)) := +theorem nth_succ_iterate (n : nat) (f : α → α) (a : α) : nth (succ n) (iterate f a) = nth n (iterate f (f a)) := by rw [nth_succ, tail_iterate] section bisim @@ -222,7 +222,7 @@ section bisim def is_bisimulation := ∀ ⦃s₁ s₂⦄, s₁ ~ s₂ → head s₁ = head s₂ ∧ tail s₁ ~ tail s₂ - lemma nth_of_bisim (bisim : is_bisimulation R) : ∀ {s₁ s₂} n, s₁ ~ s₂ → nth n s₁ = nth n s₂ ∧ drop (n+1) s₁ ~ drop (n+1) s₂ + theorem nth_of_bisim (bisim : is_bisimulation R) : ∀ {s₁ s₂} n, s₁ ~ s₂ → nth n s₁ = nth n s₂ ∧ drop (n+1) s₁ ~ drop (n+1) s₂ | s₁ s₂ 0 h := bisim h | s₁ s₂ (n+1) h := match bisim h with @@ -230,11 +230,11 @@ section bisim end -- If two streams are bisimilar, then they are equal - lemma eq_of_bisim (bisim : is_bisimulation R) : ∀ {s₁ s₂}, s₁ ~ s₂ → s₁ = s₂ := + theorem eq_of_bisim (bisim : is_bisimulation R) : ∀ {s₁ s₂}, s₁ ~ s₂ → s₁ = s₂ := λ s₁ s₂ r, stream.ext (λ n, and.elim_left (nth_of_bisim R bisim n r)) end bisim -lemma bisim_simple (s₁ s₂ : stream α) : head s₁ = head s₂ → s₁ = tail s₁ → s₂ = tail s₂ → s₁ = s₂ := +theorem bisim_simple (s₁ s₂ : stream α) : head s₁ = head s₂ → s₁ = tail s₁ → s₂ = tail s₂ → s₁ = s₂ := assume hh ht₁ ht₂, eq_of_bisim (λ s₁ s₂, head s₁ = head s₂ ∧ s₁ = tail s₁ ∧ s₂ = tail s₂) (λ s₁ s₂ ⟨h₁, h₂, h₃⟩, @@ -243,7 +243,7 @@ assume hh ht₁ ht₂, eq_of_bisim end) (and.intro hh (and.intro ht₁ ht₂)) -lemma coinduction {s₁ s₂ : stream α} : +theorem coinduction {s₁ s₂ : stream α} : head s₁ = head s₂ → (∀ (β : Type u) (fr : stream α → β), fr s₁ = fr s₂ → fr (tail s₁) = fr (tail s₂)) → s₁ = s₂ := assume hh ht, eq_of_bisim @@ -256,13 +256,13 @@ assume hh ht, and.intro h₁ (and.intro h₂ h₃)) (and.intro hh ht) -lemma iterate_id (a : α) : iterate id a = const a := +theorem iterate_id (a : α) : iterate id a = const a := coinduction rfl (λ β fr ch, begin rw [tail_iterate, tail_const], exact ch end) local attribute [reducible] stream -lemma map_iterate (f : α → α) (a : α) : iterate f (f a) = map f (iterate f a) := +theorem map_iterate (f : α → α) (a : α) : iterate f (f a) = map f (iterate f a) := begin apply funext, intro n, induction n with n' ih, @@ -279,23 +279,23 @@ def corec (f : α → β) (g : α → α) : α → stream β := def corec_on (a : α) (f : α → β) (g : α → α) : stream β := corec f g a -lemma corec_def (f : α → β) (g : α → α) (a : α) : corec f g a = map f (iterate g a) := +theorem corec_def (f : α → β) (g : α → α) (a : α) : corec f g a = map f (iterate g a) := rfl -lemma corec_eq (f : α → β) (g : α → α) (a : α) : corec f g a = f a :: corec f g (g a) := +theorem corec_eq (f : α → β) (g : α → α) (a : α) : corec f g a = f a :: corec f g (g a) := begin rw [corec_def, map_eq, head_iterate, tail_iterate], reflexivity end -lemma corec_id_id_eq_const (a : α) : corec id id a = const a := +theorem corec_id_id_eq_const (a : α) : corec id id a = const a := by rw [corec_def, map_id, iterate_id] -lemma corec_id_f_eq_iterate (f : α → α) (a : α) : corec id f a = iterate f a := +theorem corec_id_f_eq_iterate (f : α → α) (a : α) : corec id f a = iterate f a := rfl end corec section corec' def corec' (f : α → β × α) : α → stream β := corec (prod.fst ∘ f) (prod.snd ∘ f) -lemma corec'_eq (f : α → β × α) (a : α) : corec' f a = (f a).1 :: corec' f (f a).2 := +theorem corec'_eq (f : α → β × α) (a : α) : corec' f a = (f a).1 :: corec' f (f a).2 := corec_eq _ _ _ end corec' @@ -304,17 +304,17 @@ end corec' def unfolds (g : α → β) (f : α → α) (a : α) : stream β := corec g f a -lemma unfolds_eq (g : α → β) (f : α → α) (a : α) : unfolds g f a = g a :: unfolds g f (f a) := +theorem unfolds_eq (g : α → β) (f : α → α) (a : α) : unfolds g f a = g a :: unfolds g f (f a) := begin unfold unfolds, rw [corec_eq] end -lemma nth_unfolds_head_tail : ∀ (n : nat) (s : stream α), nth n (unfolds head tail s) = nth n s := +theorem nth_unfolds_head_tail : ∀ (n : nat) (s : stream α), nth n (unfolds head tail s) = nth n s := begin intro n, induction n with n' ih, {intro s, reflexivity}, {intro s, rw [nth_succ, nth_succ, unfolds_eq, tail_cons, ih]} end -lemma unfolds_head_eq : ∀ (s : stream α), unfolds head tail s = s := +theorem unfolds_head_eq : ∀ (s : stream α), unfolds head tail s = s := λ s, stream.ext (λ n, nth_unfolds_head_tail n s) def interleave (s₁ s₂ : stream α) : stream α := @@ -324,18 +324,18 @@ corec_on (s₁, s₂) infix `⋈`:65 := interleave -lemma interleave_eq (s₁ s₂ : stream α) : s₁ ⋈ s₂ = head s₁ :: head s₂ :: (tail s₁ ⋈ tail s₂) := +theorem interleave_eq (s₁ s₂ : stream α) : s₁ ⋈ s₂ = head s₁ :: head s₂ :: (tail s₁ ⋈ tail s₂) := begin unfold interleave corec_on, rw corec_eq, dsimp, rw corec_eq, reflexivity end -lemma tail_interleave (s₁ s₂ : stream α) : tail (s₁ ⋈ s₂) = s₂ ⋈ (tail s₁) := +theorem tail_interleave (s₁ s₂ : stream α) : tail (s₁ ⋈ s₂) = s₂ ⋈ (tail s₁) := begin unfold interleave corec_on, rw corec_eq, reflexivity end -lemma interleave_tail_tail (s₁ s₂ : stream α) : tail s₁ ⋈ tail s₂ = tail (tail (s₁ ⋈ s₂)) := +theorem interleave_tail_tail (s₁ s₂ : stream α) : tail s₁ ⋈ tail s₂ = tail (tail (s₁ ⋈ s₂)) := begin rw [interleave_eq s₁ s₂], reflexivity end -lemma nth_interleave_left : ∀ (n : nat) (s₁ s₂ : stream α), nth (2*n) (s₁ ⋈ s₂) = nth n s₁ +theorem nth_interleave_left : ∀ (n : nat) (s₁ s₂ : stream α), nth (2*n) (s₁ ⋈ s₂) = nth n s₁ | 0 s₁ s₂ := rfl | (succ n) s₁ s₂ := begin @@ -344,7 +344,7 @@ lemma nth_interleave_left : ∀ (n : nat) (s₁ s₂ : stream α), nth (2*n) (s reflexivity end -lemma nth_interleave_right : ∀ (n : nat) (s₁ s₂ : stream α), nth (2*n+1) (s₁ ⋈ s₂) = nth n s₂ +theorem nth_interleave_right : ∀ (n : nat) (s₁ s₂ : stream α), nth (2*n+1) (s₁ ⋈ s₂) = nth n s₂ | 0 s₁ s₂ := rfl | (succ n) s₁ s₂ := begin @@ -353,11 +353,11 @@ lemma nth_interleave_right : ∀ (n : nat) (s₁ s₂ : stream α), nth (2*n+1) reflexivity end -lemma mem_interleave_left {a : α} {s₁ : stream α} (s₂ : stream α) : a ∈ s₁ → a ∈ s₁ ⋈ s₂ := +theorem mem_interleave_left {a : α} {s₁ : stream α} (s₂ : stream α) : a ∈ s₁ → a ∈ s₁ ⋈ s₂ := assume ⟨n, h⟩, exists.intro (2*n) (by rw [h, nth_interleave_left]) -lemma mem_interleave_right {a : α} {s₁ : stream α} (s₂ : stream α) : a ∈ s₂ → a ∈ s₁ ⋈ s₂ := +theorem mem_interleave_right {a : α} {s₁ : stream α} (s₂ : stream α) : a ∈ s₂ → a ∈ s₁ ⋈ s₂ := assume ⟨n, h⟩, exists.intro (2*n+1) (by rw [h, nth_interleave_right]) @@ -370,22 +370,22 @@ corec def odd (s : stream α) : stream α := even (tail s) -lemma odd_eq (s : stream α) : odd s = even (tail s) := +theorem odd_eq (s : stream α) : odd s = even (tail s) := rfl -lemma head_even (s : stream α) : head (even s) = head s := +theorem head_even (s : stream α) : head (even s) = head s := rfl -lemma tail_even (s : stream α) : tail (even s) = even (tail (tail s)) := +theorem tail_even (s : stream α) : tail (even s) = even (tail (tail s)) := begin unfold even, rw corec_eq, reflexivity end -lemma even_cons_cons (a₁ a₂ : α) (s : stream α) : even (a₁ :: a₂ :: s) = a₁ :: even s := +theorem even_cons_cons (a₁ a₂ : α) (s : stream α) : even (a₁ :: a₂ :: s) = a₁ :: even s := begin unfold even, rw corec_eq, reflexivity end -lemma even_tail (s : stream α) : even (tail s) = odd s := +theorem even_tail (s : stream α) : even (tail s) = odd s := rfl -lemma even_interleave (s₁ s₂ : stream α) : even (s₁ ⋈ s₂) = s₁ := +theorem even_interleave (s₁ s₂ : stream α) : even (s₁ ⋈ s₂) = s₁ := eq_of_bisim (λ s₁' s₁, ∃ s₂, s₁' = even (s₁ ⋈ s₂)) (λ s₁' s₁ ⟨s₂, h₁⟩, @@ -397,7 +397,7 @@ eq_of_bisim end) (exists.intro s₂ rfl) -lemma interleave_even_odd (s₁ : stream α) : even s₁ ⋈ odd s₁ = s₁ := +theorem interleave_even_odd (s₁ : stream α) : even s₁ ⋈ odd s₁ = s₁ := eq_of_bisim (λ s' s, s' = even s ⋈ odd s) (λ s' s (h : s' = even s ⋈ odd s), @@ -408,7 +408,7 @@ eq_of_bisim end) rfl -lemma nth_even : ∀ (n : nat) (s : stream α), nth n (even s) = nth (2*n) s +theorem nth_even : ∀ (n : nat) (s : stream α), nth n (even s) = nth (2*n) s | 0 s := rfl | (succ n) s := begin @@ -416,14 +416,14 @@ lemma nth_even : ∀ (n : nat) (s : stream α), nth n (even s) = nth (2*n) s rw [nth_succ, nth_succ, tail_even, nth_even], reflexivity end -lemma nth_odd : ∀ (n : nat) (s : stream α), nth n (odd s) = nth (2*n + 1) s := +theorem nth_odd : ∀ (n : nat) (s : stream α), nth n (odd s) = nth (2*n + 1) s := λ n s, begin rw [odd_eq, nth_even], reflexivity end -lemma mem_of_mem_even (a : α) (s : stream α) : a ∈ even s → a ∈ s := +theorem mem_of_mem_even (a : α) (s : stream α) : a ∈ even s → a ∈ s := assume ⟨n, h⟩, exists.intro (2*n) (by rw [h, nth_even]) -lemma mem_of_mem_odd (a : α) (s : stream α) : a ∈ odd s → a ∈ s := +theorem mem_of_mem_odd (a : α) (s : stream α) : a ∈ odd s → a ∈ s := assume ⟨n, h⟩, exists.intro (2*n+1) (by rw [h, nth_odd]) @@ -431,36 +431,36 @@ def append_stream : list α → stream α → stream α | [] s := s | (list.cons a l) s := a :: append_stream l s -lemma nil_append_stream (s : stream α) : append_stream [] s = s := +theorem nil_append_stream (s : stream α) : append_stream [] s = s := rfl -lemma cons_append_stream (a : α) (l : list α) (s : stream α) : append_stream (a::l) s = a :: append_stream l s := +theorem cons_append_stream (a : α) (l : list α) (s : stream α) : append_stream (a::l) s = a :: append_stream l s := rfl infix `++ₛ`:65 := append_stream -lemma append_append_stream : ∀ (l₁ l₂ : list α) (s : stream α), (l₁ ++ l₂) ++ₛ s = l₁ ++ₛ (l₂ ++ₛ s) +theorem append_append_stream : ∀ (l₁ l₂ : list α) (s : stream α), (l₁ ++ l₂) ++ₛ s = l₁ ++ₛ (l₂ ++ₛ s) | [] l₂ s := rfl | (list.cons a l₁) l₂ s := by rw [list.cons_append, cons_append_stream, cons_append_stream, append_append_stream] -lemma map_append_stream (f : α → β) : ∀ (l : list α) (s : stream α), map f (l ++ₛ s) = list.map f l ++ₛ map f s +theorem map_append_stream (f : α → β) : ∀ (l : list α) (s : stream α), map f (l ++ₛ s) = list.map f l ++ₛ map f s | [] s := rfl | (list.cons a l) s := by rw [cons_append_stream, list.map_cons, map_cons, cons_append_stream, map_append_stream] -lemma drop_append_stream : ∀ (l : list α) (s : stream α), drop l.length (l ++ₛ s) = s +theorem drop_append_stream : ∀ (l : list α) (s : stream α), drop l.length (l ++ₛ s) = s | [] s := by reflexivity | (list.cons a l) s := by rw [list.length_cons, add_one_eq_succ, drop_succ, cons_append_stream, tail_cons, drop_append_stream] -lemma append_stream_head_tail (s : stream α) : [head s] ++ₛ tail s = s := +theorem append_stream_head_tail (s : stream α) : [head s] ++ₛ tail s = s := by rw [cons_append_stream, nil_append_stream, stream.eta] -lemma mem_append_stream_right : ∀ {a : α} (l : list α) {s : stream α}, a ∈ s → a ∈ l ++ₛ s +theorem mem_append_stream_right : ∀ {a : α} (l : list α) {s : stream α}, a ∈ s → a ∈ l ++ₛ s | a [] s h := h | a (list.cons b l) s h := have ih : a ∈ l ++ₛ s, from mem_append_stream_right l h, mem_cons_of_mem _ ih -lemma mem_append_stream_left : ∀ {a : α} {l : list α} (s : stream α), a ∈ l → a ∈ l ++ₛ s +theorem mem_append_stream_left : ∀ {a : α} {l : list α} (s : stream α), a ∈ l → a ∈ l ++ₛ s | a [] s h := absurd h (list.not_mem_nil _) | a (list.cons b l) s h := or.elim (list.eq_or_mem_of_mem_cons h) @@ -471,17 +471,17 @@ def approx : nat → stream α → list α | 0 s := [] | (n+1) s := list.cons (head s) (approx n (tail s)) -lemma approx_zero (s : stream α) : approx 0 s = [] := +theorem approx_zero (s : stream α) : approx 0 s = [] := rfl -lemma approx_succ (n : nat) (s : stream α) : approx (succ n) s = head s :: approx n (tail s) := +theorem approx_succ (n : nat) (s : stream α) : approx (succ n) s = head s :: approx n (tail s) := rfl -lemma nth_approx : ∀ (n : nat) (s : stream α), list.nth (approx (succ n) s) n = some (nth n s) +theorem nth_approx : ∀ (n : nat) (s : stream α), list.nth (approx (succ n) s) n = some (nth n s) | 0 s := rfl | (n+1) s := begin rw [approx_succ, add_one_eq_succ, list.nth_succ, nth_approx], reflexivity end -lemma append_approx_drop : ∀ (n : nat) (s : stream α), append_stream (approx n s) (drop n s) = s := +theorem append_approx_drop : ∀ (n : nat) (s : stream α), append_stream (approx n s) (drop n s) = s := begin intro n, induction n with n' ih, @@ -489,9 +489,9 @@ begin {intro s, rw [approx_succ, drop_succ, cons_append_stream, ih (tail s), stream.eta]} end --- Take lemma reduces a proof of equality of infinite streams to an +-- Take theorem reduces a proof of equality of infinite streams to an -- induction over all their finite approximations. -lemma take_lemma (s₁ s₂ : stream α) : (∀ (n : nat), approx n s₁ = approx n s₂) → s₁ = s₂ := +theorem take_theorem (s₁ s₂ : stream α) : (∀ (n : nat), approx n s₁ = approx n s₂) → s₁ = s₂ := begin intro h, apply stream.ext, intro n, induction n with n ih, @@ -518,7 +518,7 @@ def cycle : Π (l : list α), l ≠ [] → stream α | [] h := absurd rfl h | (list.cons a l) h := corec cycle_f cycle_g (a, l, a, l) -lemma cycle_eq : ∀ (l : list α) (h : l ≠ []), cycle l h = l ++ₛ cycle l h +theorem cycle_eq : ∀ (l : list α) (h : l ≠ []), cycle l h = l ++ₛ cycle l h | [] h := absurd rfl h | (list.cons a l) h := have gen : ∀ l' a', corec cycle_f cycle_g (a', l', a, l) = (a' :: l') ++ₛ corec cycle_f cycle_g (a, l, a, l), @@ -530,10 +530,10 @@ lemma cycle_eq : ∀ (l : list α) (h : l ≠ []), cycle l h = l ++ₛ cycle l h end, gen l a -lemma mem_cycle {a : α} {l : list α} : ∀ (h : l ≠ []), a ∈ l → a ∈ cycle l h := +theorem mem_cycle {a : α} {l : list α} : ∀ (h : l ≠ []), a ∈ l → a ∈ cycle l h := assume h ainl, begin rw [cycle_eq], exact mem_append_stream_left _ ainl end -lemma cycle_singleton (a : α) (h : [a] ≠ []) : cycle [a] h = const a := +theorem cycle_singleton (a : α) (h : [a] ≠ []) : cycle [a] h = const a := coinduction rfl (λ β fr ch, by rwa [cycle_eq, const_eq]) @@ -541,17 +541,17 @@ coinduction def tails (s : stream α) : stream (stream α) := corec id tail (tail s) -lemma tails_eq (s : stream α) : tails s = tail s :: tails (tail s) := +theorem tails_eq (s : stream α) : tails s = tail s :: tails (tail s) := by unfold tails; rw [corec_eq]; reflexivity -lemma nth_tails : ∀ (n : nat) (s : stream α), nth n (tails s) = drop n (tail s) := +theorem nth_tails : ∀ (n : nat) (s : stream α), nth n (tails s) = drop n (tail s) := begin intro n, induction n with n' ih, {intros, reflexivity}, {intro s, rw [nth_succ, drop_succ, tails_eq, tail_cons, ih]} end -lemma tails_eq_iterate (s : stream α) : tails s = iterate tail (tail s) := +theorem tails_eq_iterate (s : stream α) : tails s = iterate tail (tail s) := rfl def inits_core (l : list α) (s : stream α) : stream (list α) := @@ -563,16 +563,16 @@ corec_on (l, s) def inits (s : stream α) : stream (list α) := inits_core [head s] (tail s) -lemma inits_core_eq (l : list α) (s : stream α) : inits_core l s = l :: inits_core (l ++ [head s]) (tail s) := +theorem inits_core_eq (l : list α) (s : stream α) : inits_core l s = l :: inits_core (l ++ [head s]) (tail s) := begin unfold inits_core corec_on, rw [corec_eq], reflexivity end -lemma tail_inits (s : stream α) : tail (inits s) = inits_core [head s, head (tail s)] (tail (tail s)) := +theorem tail_inits (s : stream α) : tail (inits s) = inits_core [head s, head (tail s)] (tail (tail s)) := begin unfold inits, rw inits_core_eq, reflexivity end -lemma inits_tail (s : stream α) : inits (tail s) = inits_core [head (tail s)] (tail (tail s)) := +theorem inits_tail (s : stream α) : inits (tail s) = inits_core [head (tail s)] (tail (tail s)) := rfl -lemma cons_nth_inits_core : ∀ (a : α) (n : nat) (l : list α) (s : stream α), +theorem cons_nth_inits_core : ∀ (a : α) (n : nat) (l : list α) (s : stream α), a :: nth n (inits_core l s) = nth n (inits_core (a::l) s) := begin intros a n, @@ -581,14 +581,14 @@ begin {intros l s, rw [nth_succ, inits_core_eq, tail_cons, ih, inits_core_eq (a::l) s], reflexivity } end -lemma nth_inits : ∀ (n : nat) (s : stream α), nth n (inits s) = approx (succ n) s := +theorem nth_inits : ∀ (n : nat) (s : stream α), nth n (inits s) = approx (succ n) s := begin intro n, induction n with n' ih, {intros, reflexivity}, {intros, rw [nth_succ, approx_succ, ← ih, tail_inits, inits_tail, cons_nth_inits_core]} end -lemma inits_eq (s : stream α) : inits s = [head s] :: map (list.cons (head s)) (inits (tail s)) := +theorem inits_eq (s : stream α) : inits s = [head s] :: map (list.cons (head s)) (inits (tail s)) := begin apply stream.ext, intro n, cases n, @@ -596,7 +596,7 @@ begin {rw [nth_inits, nth_succ, tail_cons, nth_map, nth_inits], reflexivity} end -lemma zip_inits_tails (s : stream α) : zip append_stream (inits s) (tails s) = const s := +theorem zip_inits_tails (s : stream α) : zip append_stream (inits s) (tails s) = const s := begin apply stream.ext, intro n, rw [nth_zip, nth_inits, nth_tails, nth_const, approx_succ, @@ -611,24 +611,24 @@ def apply (f : stream (α → β)) (s : stream α) : stream β := infix `⊛`:75 := apply -- input as \o* -lemma identity (s : stream α) : pure id ⊛ s = s := +theorem identity (s : stream α) : pure id ⊛ s = s := rfl -lemma composition (g : stream (β → δ)) (f : stream (α → β)) (s : stream α) : pure comp ⊛ g ⊛ f ⊛ s = g ⊛ (f ⊛ s) := +theorem composition (g : stream (β → δ)) (f : stream (α → β)) (s : stream α) : pure comp ⊛ g ⊛ f ⊛ s = g ⊛ (f ⊛ s) := rfl -lemma homomorphism (f : α → β) (a : α) : pure f ⊛ pure a = pure (f a) := +theorem homomorphism (f : α → β) (a : α) : pure f ⊛ pure a = pure (f a) := rfl -lemma interchange (fs : stream (α → β)) (a : α) : fs ⊛ pure a = pure (λ f : α → β, f a) ⊛ fs := +theorem interchange (fs : stream (α → β)) (a : α) : fs ⊛ pure a = pure (λ f : α → β, f a) ⊛ fs := rfl -lemma map_eq_apply (f : α → β) (s : stream α) : map f s = pure f ⊛ s := +theorem map_eq_apply (f : α → β) (s : stream α) : map f s = pure f ⊛ s := rfl def nats : stream nat := λ n, n -lemma nth_nats (n : nat) : nth n nats = n := +theorem nth_nats (n : nat) : nth n nats = n := rfl -lemma nats_eq : nats = 0 :: map succ nats := +theorem nats_eq : nats = 0 :: map succ nats := begin apply stream.ext, intro n, cases n, reflexivity, rw [nth_succ], reflexivity diff --git a/data/vector.lean b/data/vector.lean index b835dd5bc6b1b..ff908d73ddabc 100644 --- a/data/vector.lean +++ b/data/vector.lean @@ -65,9 +65,9 @@ def append {n m : nat} : vector α n → vector α m → vector α (n + m) def map (f : α → β) : vector α n → vector β n | ⟨ l, h ⟩ := ⟨ list.map f l, by simp * ⟩ -@[simp] lemma map_nil (f : α → β) : map f nil = nil := rfl +@[simp] theorem map_nil (f : α → β) : map f nil = nil := rfl -lemma map_cons (f : α → β) (a : α) : Π (v : vector α n), map f (a::v) = f a :: map f v +theorem map_cons (f : α → β) (a : α) : Π (v : vector α n), map f (a::v) = f a :: map f v | ⟨l,h⟩ := rfl def map₂ (f : α → β → φ) : vector α n → vector β n → vector φ n @@ -106,30 +106,30 @@ section accum end accum -protected lemma eq {n : ℕ} : ∀ (a1 a2 : vector α n), to_list a1 = to_list a2 → a1 = a2 +protected theorem eq {n : ℕ} : ∀ (a1 a2 : vector α n), to_list a1 = to_list a2 → a1 = a2 | ⟨x, h1⟩ ⟨._, h2⟩ rfl := rfl protected theorem eq_nil (v : vector α 0) : v = [] := v.eq [] (list.eq_nil_of_length_eq_zero v.2) -@[simp] lemma to_list_mk (v : list α) (P : list.length v = n) : to_list (subtype.mk v P) = v := +@[simp] theorem to_list_mk (v : list α) (P : list.length v = n) : to_list (subtype.mk v P) = v := rfl -@[simp] lemma to_list_nil : to_list [] = @list.nil α := +@[simp] theorem to_list_nil : to_list [] = @list.nil α := rfl -@[simp] lemma to_list_length (v : vector α n) : (to_list v).length = n := v.2 +@[simp] theorem to_list_length (v : vector α n) : (to_list v).length = n := v.2 -@[simp] lemma to_list_cons (a : α) (v : vector α n) : to_list (a :: v) = a :: to_list v := +@[simp] theorem to_list_cons (a : α) (v : vector α n) : to_list (a :: v) = a :: to_list v := begin cases v, reflexivity end -@[simp] lemma to_list_append {n m : nat} (v : vector α n) (w : vector α m) : to_list (append v w) = to_list v ++ to_list w := +@[simp] theorem to_list_append {n m : nat} (v : vector α n) (w : vector α m) : to_list (append v w) = to_list v ++ to_list w := begin cases v, cases w, reflexivity end -@[simp] lemma to_list_drop {n m : ℕ} (v : vector α m) : to_list (drop n v) = list.drop n (to_list v) := +@[simp] theorem to_list_drop {n m : ℕ} (v : vector α m) : to_list (drop n v) = list.drop n (to_list v) := begin cases v, reflexivity end -@[simp] lemma to_list_take {n m : ℕ} (v : vector α m) : to_list (take n v) = list.take n (to_list v) := +@[simp] theorem to_list_take {n m : ℕ} (v : vector α m) : to_list (take n v) = list.take n (to_list v) := begin cases v, reflexivity end end vector diff --git a/logic/basic.lean b/logic/basic.lean index dca1cf50712cf..8d90c70cd0757 100644 --- a/logic/basic.lean +++ b/logic/basic.lean @@ -21,26 +21,26 @@ section miscellany universes u v variables {α : Type u} {β : Type v} -lemma eq_iff_le_and_le {α : Type u} [weak_order α] {a b : α} : a = b ↔ (a ≤ b ∧ b ≤ a) := +theorem eq_iff_le_and_le {α : Type u} [weak_order α] {a b : α} : a = b ↔ (a ≤ b ∧ b ≤ a) := ⟨assume eq, eq ▸ ⟨le_refl a, le_refl a⟩, assume ⟨ab, ba⟩, le_antisymm ab ba⟩ @[simp] -lemma prod.mk.inj_iff {α : Type u} {β : Type v} {a₁ a₂ : α} {b₁ b₂ : β} : +theorem prod.mk.inj_iff {α : Type u} {β : Type v} {a₁ a₂ : α} {b₁ b₂ : β} : (a₁, b₁) = (a₂, b₂) ↔ (a₁ = a₂ ∧ b₁ = b₂) := ⟨prod.mk.inj, by cc⟩ @[simp] -lemma prod.forall {α : Type u} {β : Type v} {p : α × β → Prop} : +theorem prod.forall {α : Type u} {β : Type v} {p : α × β → Prop} : (∀x, p x) ↔ (∀a b, p (a, b)) := ⟨assume h a b, h (a, b), assume h ⟨a, b⟩, h a b⟩ @[simp] -lemma prod.exists {α : Type u} {β : Type v} {p : α × β → Prop} : +theorem prod.exists {α : Type u} {β : Type v} {p : α × β → Prop} : (∃x, p x) ↔ (∃a b, p (a, b)) := ⟨assume ⟨⟨a, b⟩, h⟩, ⟨a, b, h⟩, assume ⟨a, b, h⟩, ⟨⟨a, b⟩, h⟩⟩ @[simp] -lemma set_of_subset_set_of {p q : α → Prop} : {a | p a} ⊆ {a | q a} = (∀a, p a → q a) := +theorem set_of_subset_set_of {p q : α → Prop} : {a | p a} ⊆ {a | q a} = (∀a, p a → q a) := rfl @@ -145,13 +145,13 @@ iff.intro theorem or_iff_or (h1 : a ↔ c) (h2 : b ↔ d) : (a ∨ b) ↔ (c ∨ d) := iff.intro (or.imp (iff.mp h1) (iff.mp h2)) (or.imp (iff.mpr h1) (iff.mpr h2)) -lemma or_of_not_implies {a b : Prop} [decidable a] (h : ¬ a → b) : a ∨ b := +theorem or_of_not_implies {a b : Prop} [decidable a] (h : ¬ a → b) : a ∨ b := dite _ or.inl (or.inr ∘ h) -lemma or_of_not_implies' {a b : Prop} [decidable b] (h : ¬ b → a) : a ∨ b := +theorem or_of_not_implies' {a b : Prop} [decidable b] (h : ¬ b → a) : a ∨ b := dite _ or.inr (or.inl ∘ h) -lemma not_imp_iff_not_imp {a b : Prop} [decidable a] : +theorem not_imp_iff_not_imp {a b : Prop} [decidable a] : (¬ a → ¬ b) ↔ (b → a) := ⟨assume h hb, by_contradiction $ assume na, h na hb, mt⟩ @@ -242,7 +242,7 @@ by rewrite [←not_and_iff, not_not_iff] /- other identities -/ -lemma or_imp_iff_and_imp {a b c : Prop} : ((a ∨ b) → c) ↔ ((a → c) ∧ (b → c)) := +theorem or_imp_iff_and_imp {a b c : Prop} : ((a ∨ b) → c) ↔ ((a → c) ∧ (b → c)) := ⟨assume h, ⟨assume ha, h (or.inl ha), assume hb, h (or.inr hb)⟩, assume ⟨ha, hb⟩, or.rec ha hb⟩ @@ -321,7 +321,7 @@ iff.intro (∃x, q ∧ p x) ↔ q ∧ (∃x, p x) := ⟨assume ⟨x, hq, hp⟩, ⟨hq, x, hp⟩, assume ⟨hq, x, hp⟩, ⟨x, hq, hp⟩⟩ -lemma forall_eq {α : Type u} {p : α → Prop} {a' : α} : (∀a, a = a' → p a) ↔ p a' := +theorem forall_eq {α : Type u} {p : α → Prop} {a' : α} : (∀a, a = a' → p a) ↔ p a' := ⟨assume h, h a' rfl, assume h a eq, eq.symm ▸ h⟩ theorem forall_or_distrib_left {q : Prop} {p : α → Prop} [decidable q] : diff --git a/tactic/converter/binders.lean b/tactic/converter/binders.lean index d71bbbc1ec1bf..a0b42d14d7f35 100644 --- a/tactic/converter/binders.lean +++ b/tactic/converter/binders.lean @@ -137,15 +137,15 @@ meta def binder_eq_elim.old_conv (b : binder_eq_elim) : old_conv unit := do b.check x (bd.instantiate_var x), b.adapt_rel b.push -lemma {u v} exists_comm {α : Sort u} {β : Sort v} (p : α → β → Prop) : +theorem {u v} exists_comm {α : Sort u} {β : Sort v} (p : α → β → Prop) : (∃a b, p a b) ↔ (∃b a, p a b) := ⟨λ⟨a, ⟨b, h⟩⟩, ⟨b, ⟨a, h⟩⟩, λ⟨a, ⟨b, h⟩⟩, ⟨b, ⟨a, h⟩⟩⟩ -lemma {u v} exists_elim_eq_left {α : Sort u} (a : α) (p : Π(a':α), a' = a → Prop) : +theorem {u v} exists_elim_eq_left {α : Sort u} (a : α) (p : Π(a':α), a' = a → Prop) : (∃(a':α)(h : a' = a), p a' h) ↔ p a rfl := ⟨λ⟨a', ⟨h, p_h⟩⟩, match a', h, p_h with ._, rfl, h := h end, λh, ⟨a, rfl, h⟩⟩ -lemma {u v} exists_elim_eq_right {α : Sort u} (a : α) (p : Π(a':α), a = a' → Prop) : +theorem {u v} exists_elim_eq_right {α : Sort u} (a : α) (p : Π(a':α), a = a' → Prop) : (∃(a':α)(h : a = a'), p a' h) ↔ p a rfl := ⟨λ⟨a', ⟨h, p_h⟩⟩, match a', h, p_h with ._, rfl, h := h end, λh, ⟨a, rfl, h⟩⟩ @@ -156,15 +156,15 @@ meta def exists_eq_elim : binder_eq_elim := apply_congr := congr_binder ``exists_congr, apply_elim_eq := apply' ``exists_elim_eq_left <|> apply' ``exists_elim_eq_right } -lemma {u v} forall_comm {α : Sort u} {β : Sort v} (p : α → β → Prop) : +theorem {u v} forall_comm {α : Sort u} {β : Sort v} (p : α → β → Prop) : (∀a b, p a b) ↔ (∀b a, p a b) := ⟨assume h b a, h a b, assume h b a, h a b⟩ -lemma {u v} forall_elim_eq_left {α : Sort u} (a : α) (p : Π(a':α), a' = a → Prop) : +theorem {u v} forall_elim_eq_left {α : Sort u} (a : α) (p : Π(a':α), a' = a → Prop) : (∀(a':α)(h : a' = a), p a' h) ↔ p a rfl := ⟨λh, h a rfl, λh a' h_eq, match a', h_eq with ._, rfl := h end⟩ -lemma {u v} forall_elim_eq_right {α : Sort u} (a : α) (p : Π(a':α), a = a' → Prop) : +theorem {u v} forall_elim_eq_right {α : Sort u} (a : α) (p : Π(a':α), a = a' → Prop) : (∀(a':α)(h : a = a'), p a' h) ↔ p a rfl := ⟨λh, h a rfl, λh a' h_eq, match a', h_eq with ._, rfl := h end⟩ @@ -194,20 +194,20 @@ universes u v w w₂ variables {α : Type u} {β : Type v} {ι : Sort w} {ι₂ : Sort w₂} {s t : set α} {a : α} @[simp] -lemma mem_image {f : α → β} {b : β} : b ∈ set.image f s = ∃a, a ∈ s ∧ f a = b := +theorem mem_image {f : α → β} {b : β} : b ∈ set.image f s = ∃a, a ∈ s ∧ f a = b := rfl section open lattice variables [complete_lattice α] -lemma Inf_image {s : set β} {f : β → α} : Inf (set.image f s) = (⨅ a ∈ s, f a) := +theorem Inf_image {s : set β} {f : β → α} : Inf (set.image f s) = (⨅ a ∈ s, f a) := begin simp [Inf_eq_infi, infi_and], conversion infi_eq_elim.old_conv, end -lemma Sup_image {s : set β} {f : β → α} : Sup (set.image f s) = (⨆ a ∈ s, f a) := +theorem Sup_image {s : set β} {f : β → α} : Sup (set.image f s) = (⨆ a ∈ s, f a) := begin simp [Sup_eq_supr, supr_and], conversion supr_eq_elim.old_conv, diff --git a/tests/examples.lean b/tests/examples.lean index 573469a11c299..f8868b336ea9a 100644 --- a/tests/examples.lean +++ b/tests/examples.lean @@ -41,7 +41,7 @@ begin end @[simp] -lemma mem_set_of {a : α} {p : α → Prop} : a ∈ {a | p a} = p a := rfl +theorem mem_set_of {a : α} {p : α → Prop} : a ∈ {a | p a} = p a := rfl -- TODO: write a tactic to unfold specific instances of generic notation? theorem subset_def {s t : set α} : (s ⊆ t) = ∀ x, x ∈ s → x ∈ t := rfl diff --git a/tests/finish2.lean b/tests/finish2.lean index 0a588f90e47cf..de9e13ef19254 100644 --- a/tests/finish2.lean +++ b/tests/finish2.lean @@ -215,7 +215,7 @@ example (n2 : ¬A₂) (n3 : ¬A₃) (b1 : B₁) (b2 : B₂) (b3 : B₃) (b4 : B section club variables Scottish RedSocks WearKilt Married GoOutSunday : Prop -lemma NoMember : (¬Scottish → RedSocks) → (WearKilt ∨ ¬RedSocks) → (Married → ¬GoOutSunday) → +theorem NoMember : (¬Scottish → RedSocks) → (WearKilt ∨ ¬RedSocks) → (Married → ¬GoOutSunday) → (GoOutSunday ↔ Scottish) → (WearKilt → Scottish ∧ Married) → (Scottish → WearKilt) → false := by finish end club diff --git a/topology/continuity.lean b/topology/continuity.lean index bec0328d2fe00..8f988e166cde8 100644 --- a/topology/continuity.lean +++ b/topology/continuity.lean @@ -14,14 +14,14 @@ variables {α : Type u} {β : Type v} {γ : Type w} {ι : Sort x} theorem classical.cases {p : Prop → Prop} (h1 : p true) (h2 : p false) : ∀a, p a := assume a, classical.cases_on a h1 h2 -lemma univ_eq_true_false : univ = ({true, false} : set Prop) := +theorem univ_eq_true_false : univ = ({true, false} : set Prop) := eq.symm $ top_unique $ classical.cases (by simp) (by simp) @[simp] -lemma false_neq_true : false ≠ true := +theorem false_neq_true : false ≠ true := begin intro h, rw [h], trivial end -lemma subtype.val_image {p : α → Prop} {s : set (subtype p)} : +theorem subtype.val_image {p : α → Prop} {s : set (subtype p)} : subtype.val '' s = {x | ∃h : p x, (⟨x, h⟩ : subtype p) ∈ s} := set.ext $ assume a, ⟨assume ⟨⟨a', ha'⟩, in_s, (h_eq : a' = a)⟩, h_eq ▸ ⟨ha', in_s⟩, @@ -32,14 +32,14 @@ variables [topological_space α] [topological_space β] [topological_space γ] def continuous (f : α → β) := ∀s, open' s → open' (vimage f s) -lemma continuous_id : continuous (id : α → α) := +theorem continuous_id : continuous (id : α → α) := assume s h, h -lemma continuous_compose {f : α → β} {g : β → γ} (hf : continuous f) (hg : continuous g): +theorem continuous_compose {f : α → β} {g : β → γ} (hf : continuous f) (hg : continuous g): continuous (g ∘ f) := assume s h, hf _ (hg s h) -lemma continuous_iff_towards {f : α → β} : +theorem continuous_iff_towards {f : α → β} : continuous f ↔ (∀x, towards f (nhds x) (nhds (f x))) := ⟨assume hf : continuous f, assume x s, show s ∈ (nhds (f x)).sets → s ∈ (map f (nhds x)).sets, @@ -62,81 +62,81 @@ open topological_space variable {f : α → β} -lemma continuous_iff_induced_le {t₁ : tspace α} {t₂ : tspace β} : +theorem continuous_iff_induced_le {t₁ : tspace α} {t₂ : tspace β} : cont t₁ t₂ f ↔ (induced f t₂ ≤ t₁) := ⟨assume hc s ⟨t, ht, s_eq⟩, s_eq.symm ▸ hc t ht, assume hle s h, hle _ ⟨_, h, rfl⟩⟩ -lemma continuous_eq_le_coinduced {t₁ : tspace α} {t₂ : tspace β} : +theorem continuous_eq_le_coinduced {t₁ : tspace α} {t₂ : tspace β} : cont t₁ t₂ f = (t₂ ≤ coinduced f t₁) := rfl -lemma continuous_generated_from {t : tspace α} {b : set (set β)} +theorem continuous_generated_from {t : tspace α} {b : set (set β)} (h : ∀s∈b, open' (vimage f s)) : cont t (generate_from b) f := assume s hs, generate_open.rec_on hs h open_univ (assume s t _ _, open_inter) (assume t _ h, by rw [vimage_sUnion]; exact (open_Union $ assume s, open_Union $ assume hs, h s hs)) -lemma continuous_induced_dom {t : tspace β} : cont (induced f t) t f := +theorem continuous_induced_dom {t : tspace β} : cont (induced f t) t f := assume s h, ⟨_, h, rfl⟩ -lemma continuous_induced_rng {g : γ → α} {t₂ : tspace β} {t₁ : tspace γ} +theorem continuous_induced_rng {g : γ → α} {t₂ : tspace β} {t₁ : tspace γ} (h : cont t₁ t₂ (f ∘ g)) : cont t₁ (induced f t₂) g := assume s ⟨t, ht, s_eq⟩, s_eq.symm ▸ h t ht -lemma continuous_coinduced_rng {t : tspace α} : cont t (coinduced f t) f := +theorem continuous_coinduced_rng {t : tspace α} : cont t (coinduced f t) f := assume s h, h -lemma continuous_coinduced_dom {g : β → γ} {t₁ : tspace α} {t₂ : tspace γ} +theorem continuous_coinduced_dom {g : β → γ} {t₁ : tspace α} {t₂ : tspace γ} (h : cont t₁ t₂ (g ∘ f)) : cont (coinduced f t₁) t₂ g := assume s hs, h s hs -lemma continuous_inf_dom {t₁ t₂ : tspace α} {t₃ : tspace β} +theorem continuous_inf_dom {t₁ t₂ : tspace α} {t₃ : tspace β} (h₁ : cont t₁ t₃ f) (h₂ : cont t₂ t₃ f) : cont (t₁ ⊓ t₂) t₃ f := assume s h, ⟨h₁ s h, h₂ s h⟩ -lemma continuous_inf_rng_left {t₁ : tspace α} {t₃ t₂ : tspace β} +theorem continuous_inf_rng_left {t₁ : tspace α} {t₃ t₂ : tspace β} (h : cont t₁ t₂ f) : cont t₁ (t₂ ⊓ t₃) f := assume s hs, h s hs.left -lemma continuous_inf_rng_right {t₁ : tspace α} {t₃ t₂ : tspace β} +theorem continuous_inf_rng_right {t₁ : tspace α} {t₃ t₂ : tspace β} (h : cont t₁ t₃ f) : cont t₁ (t₂ ⊓ t₃) f := assume s hs, h s hs.right -lemma continuous_Inf_dom {t₁ : set (tspace α)} {t₂ : tspace β} +theorem continuous_Inf_dom {t₁ : set (tspace α)} {t₂ : tspace β} (h : ∀t∈t₁, cont t t₂ f) : cont (Inf t₁) t₂ f := assume s hs t ht, h t ht s hs -lemma continuous_Inf_rng {t₁ : tspace α} {t₂ : set (tspace β)} +theorem continuous_Inf_rng {t₁ : tspace α} {t₂ : set (tspace β)} {t : tspace β} (h₁ : t ∈ t₂) (hf : cont t₁ t f) : cont t₁ (Inf t₂) f := assume s hs, hf s $ hs t h₁ -lemma continuous_infi_dom {t₁ : ι → tspace α} {t₂ : tspace β} +theorem continuous_infi_dom {t₁ : ι → tspace α} {t₂ : tspace β} (h : ∀i, cont (t₁ i) t₂ f) : cont (infi t₁) t₂ f := continuous_Inf_dom $ assume t ⟨i, (t_eq : t = t₁ i)⟩, t_eq.symm ▸ h i -lemma continuous_infi_rng {t₁ : tspace α} {t₂ : ι → tspace β} +theorem continuous_infi_rng {t₁ : tspace α} {t₂ : ι → tspace β} {i : ι} (h : cont t₁ (t₂ i) f) : cont t₁ (infi t₂) f := continuous_Inf_rng ⟨i, rfl⟩ h -lemma continuous_top {t : tspace β} : cont ⊤ t f := +theorem continuous_top {t : tspace β} : cont ⊤ t f := assume s h, trivial -lemma continuous_bot {t : tspace α} : cont t ⊥ f := +theorem continuous_bot {t : tspace α} : cont t ⊥ f := continuous_Inf_rng (mem_univ $ coinduced f t) continuous_coinduced_rng -lemma continuous_sup_rng {t₁ : tspace α} {t₂ t₃ : tspace β} +theorem continuous_sup_rng {t₁ : tspace α} {t₂ t₃ : tspace β} (h₁ : cont t₁ t₂ f) (h₂ : cont t₁ t₃ f) : cont t₁ (t₂ ⊔ t₃) f := continuous_Inf_rng (show t₂ ≤ coinduced f t₁ ∧ t₃ ≤ coinduced f t₁, from ⟨h₁, h₂⟩) continuous_coinduced_rng -lemma continuous_sup_dom_left {t₁ t₂ : tspace α} {t₃ : tspace β} +theorem continuous_sup_dom_left {t₁ t₂ : tspace α} {t₃ : tspace β} (h : cont t₁ t₃ f) : cont (t₁ ⊔ t₂) t₃ f := continuous_Inf_dom $ assume t ⟨h₁, h₂⟩ s hs, h₁ _ $ h s hs -lemma continuous_sup_dom_right {t₁ t₂ : tspace α} {t₃ : tspace β} +theorem continuous_sup_dom_right {t₁ t₂ : tspace α} {t₃ : tspace β} (h : cont t₂ t₃ f) : cont (t₁ ⊔ t₂) t₃ f := continuous_Inf_dom $ assume t ⟨h₁, h₂⟩ s hs, h₂ _ $ h s hs @@ -146,10 +146,10 @@ section sierpinski variables [topological_space α] @[simp] -lemma open_singleton_true : open' ({true} : set Prop) := +theorem open_singleton_true : open' ({true} : set Prop) := topological_space.generate_open.basic _ (by simp) -lemma continuous_Prop {p : α → Prop} : continuous p ↔ open' {x | p x} := +theorem continuous_Prop {p : α → Prop} : continuous p ↔ open' {x | p x} := ⟨assume h : continuous p, have open' (vimage p {true}), from h _ open_singleton_true, @@ -164,10 +164,10 @@ section induced open topological_space variables [t : topological_space β] {f : α → β} -lemma open_induced {s : set β} (h : open' s) : (induced f t).open' (vimage f s) := +theorem open_induced {s : set β} (h : open' s) : (induced f t).open' (vimage f s) := ⟨s, h, rfl⟩ -lemma nhds_induced_eq_vmap {a : α} : @nhds α (induced f t) a = vmap f (nhds (f a)) := +theorem nhds_induced_eq_vmap {a : α} : @nhds α (induced f t) a = vmap f (nhds (f a)) := le_antisymm (assume s ⟨s', hs', (h_s : vimage f s' ⊆ s)⟩, have ∃t':set β, open' t' ∧ t' ⊆ s' ∧ f a ∈ t', @@ -185,7 +185,7 @@ le_antisymm exact ⟨s', subset.refl _, s', open_s', subset.refl _, by rw [s_eq] at as; assumption⟩ end) -lemma map_nhds_induced_eq {a : α} (h : image f univ ∈ (nhds (f a)).sets) : +theorem map_nhds_induced_eq {a : α} (h : image f univ ∈ (nhds (f a)).sets) : map f (@nhds α (induced f t) a) = nhds (f a) := le_antisymm ((@continuous_iff_towards α β (induced f t) _ _).mp continuous_induced_dom a) @@ -206,21 +206,21 @@ section prod open topological_space variables [topological_space α] [topological_space β] [topological_space γ] -lemma continuous_fst : continuous (@prod.fst α β) := +theorem continuous_fst : continuous (@prod.fst α β) := continuous_sup_dom_left continuous_induced_dom -lemma continuous_snd : continuous (@prod.snd α β) := +theorem continuous_snd : continuous (@prod.snd α β) := continuous_sup_dom_right continuous_induced_dom -lemma continuous_prod_mk {f : γ → α} {g : γ → β} +theorem continuous_prod_mk {f : γ → α} {g : γ → β} (hf : continuous f) (hg : continuous g) : continuous (λx, prod.mk (f x) (g x)) := continuous_sup_rng (continuous_induced_rng hf) (continuous_induced_rng hg) -lemma open_set_prod {s : set α} {t : set β} (hs : open' s) (ht: open' t) : +theorem open_set_prod {s : set α} {t : set β} (hs : open' s) (ht: open' t) : open' (set.prod s t) := open_inter (continuous_fst s hs) (continuous_snd t ht) -lemma prod_eq_generate_from [tα : topological_space α] [tβ : topological_space β] : +theorem prod_eq_generate_from [tα : topological_space α] [tβ : topological_space β] : prod.topological_space = generate_from {g | ∃(s:set α) (t:set β), open' s ∧ open' t ∧ g = set.prod s t} := le_antisymm @@ -233,7 +233,7 @@ le_antisymm this ▸ (generate_open.basic _ ⟨univ, t, open_univ, ht, rfl⟩))) (generate_from_le $ assume g ⟨s, t, hs, ht, g_eq⟩, g_eq.symm ▸ open_set_prod hs ht) -lemma nhds_prod_eq {a : α} {b : β} : nhds (a, b) = filter.prod (nhds a) (nhds b) := +theorem nhds_prod_eq {a : α} {b : β} : nhds (a, b) = filter.prod (nhds a) (nhds b) := by rw [prod_eq_generate_from, nhds_generate_from]; exact le_antisymm (le_infi $ assume s, le_infi $ assume hs, le_infi $ assume t, le_infi $ assume ht, @@ -254,7 +254,7 @@ by rw [prod_eq_generate_from, nhds_generate_from]; (mem_nhds_sets_iff.mpr ⟨t', subset.refl t', ht', hb⟩) end) -lemma closure_prod_eq {s : set α} {t : set β} : +theorem closure_prod_eq {s : set α} {t : set β} : closure (set.prod s t) = set.prod (closure s) (closure t) := set.ext $ assume ⟨a, b⟩, have filter.prod (nhds a) (nhds b) ⊓ principal (set.prod s t) = @@ -267,13 +267,13 @@ end prod section sum variables [topological_space α] [topological_space β] [topological_space γ] -lemma continuous_inl : continuous (@sum.inl α β) := +theorem continuous_inl : continuous (@sum.inl α β) := continuous_inf_rng_left continuous_coinduced_rng -lemma continuous_inr : continuous (@sum.inr α β) := +theorem continuous_inr : continuous (@sum.inr α β) := continuous_inf_rng_right continuous_coinduced_rng -lemma continuous_sum_rec {f : α → γ} {g : β → γ} +theorem continuous_sum_rec {f : α → γ} {g : β → γ} (hf : continuous f) (hg : continuous g) : @continuous (α ⊕ β) γ _ _ (@sum.rec α β (λ_, γ) f g) := continuous_inf_dom hf hg @@ -282,18 +282,18 @@ end sum section subtype variables [topological_space α] [topological_space β] {p : α → Prop} -lemma continuous_subtype_val : continuous (@subtype.val α p) := +theorem continuous_subtype_val : continuous (@subtype.val α p) := continuous_induced_dom -lemma continuous_subtype_mk {f : β → α} +theorem continuous_subtype_mk {f : β → α} (hp : ∀x, p (f x)) (h : continuous f) : continuous (λx, (⟨f x, hp x⟩ : subtype p)) := continuous_induced_rng h -lemma map_nhds_subtype_val_eq {a : α} {ha : p a} (h : {a | p a} ∈ (nhds a).sets) : +theorem map_nhds_subtype_val_eq {a : α} {ha : p a} (h : {a | p a} ∈ (nhds a).sets) : map (@subtype.val α p) (nhds ⟨a, ha⟩) = nhds a := map_nhds_induced_eq (by simp [subtype.val_image, h]) -lemma continuous_subtype_nhds_cover {f : α → β} {c : ι → α → Prop} +theorem continuous_subtype_nhds_cover {f : α → β} {c : ι → α → Prop} (c_cover : ∀x, ∃i, c i x ∧ {x | c i x} ∈ (nhds x).sets) (f_cont : ∀i, continuous (λ(x : subtype (c i)), f x.val)) : continuous f := @@ -308,13 +308,13 @@ end subtype section pi -lemma nhds_pi {ι : Type u} {α : ι → Type v} [t : ∀i, topological_space (α i)] {a : Πi, α i} : +theorem nhds_pi {ι : Type u} {α : ι → Type v} [t : ∀i, topological_space (α i)] {a : Πi, α i} : nhds a = (⨅i, vmap (λx, x i) (nhds (a i))) := calc nhds a = (⨅i, @nhds _ (@topological_space.induced _ _ (λx:Πi, α i, x i) (t i)) a) : nhds_supr ... = (⨅i, vmap (λx, x i) (nhds (a i))) : by simp [nhds_induced_eq_vmap] /-- Tychonoff's theorem -/ -lemma compact_pi_infinite {ι : Type v} {α : ι → Type u} [∀i:ι, topological_space (α i)] +theorem compact_pi_infinite {ι : Type v} {α : ι → Type u} [∀i:ι, topological_space (α i)] {s : Πi:ι, set (α i)} : (∀i, compact (s i)) → compact {x : Πi:ι, α i | ∀i, x i ∈ s i} := begin simp [compact_iff_ultrafilter_le_nhds, nhds_pi], diff --git a/topology/topological_space.lean b/topology/topological_space.lean index aa5c26cad0b00..d5bf24642fe8d 100644 --- a/topology/topological_space.lean +++ b/topology/topological_space.lean @@ -22,7 +22,7 @@ section topological_space variables {α : Type u} {β : Type v} {ι : Sort w} {a a₁ a₂ : α} {s s₁ s₂ : set α} -lemma topological_space_eq {f g : topological_space α} (h' : f.open' = g.open') : f = g := +theorem topological_space_eq {f g : topological_space α} (h' : f.open' = g.open') : f = g := begin cases f with a, cases g with b, have h : a = b, assumption, @@ -38,21 +38,21 @@ include t def open' (s : set α) : Prop := topological_space.open' t s @[simp] -lemma open_univ : open' (univ : set α) := topological_space.open_univ t +theorem open_univ : open' (univ : set α) := topological_space.open_univ t -lemma open_inter (h₁ : open' s₁) (h₂ : open' s₂) : open' (s₁ ∩ s₂) := topological_space.open_inter t s₁ s₂ h₁ h₂ +theorem open_inter (h₁ : open' s₁) (h₂ : open' s₂) : open' (s₁ ∩ s₂) := topological_space.open_inter t s₁ s₂ h₁ h₂ -lemma open_sUnion {s : set (set α)} (h : ∀t ∈ s, open' t) : open' (⋃₀ s) := topological_space.open_sUnion t s h +theorem open_sUnion {s : set (set α)} (h : ∀t ∈ s, open' t) : open' (⋃₀ s) := topological_space.open_sUnion t s h end variables [topological_space α] -lemma open_Union {f : ι → set α} (h : ∀i, open' (f i)) : open' (⋃i, f i) := +theorem open_Union {f : ι → set α} (h : ∀i, open' (f i)) : open' (⋃i, f i) := open_sUnion $ assume t ⟨i, (heq : t = f i)⟩, heq.symm ▸ h i @[simp] -lemma open_empty : open' (∅ : set α) := +theorem open_empty : open' (∅ : set α) := have open' (⋃₀ ∅ : set α), from open_sUnion (assume a, false.elim), by simp at this; assumption @@ -60,76 +60,76 @@ by simp at this; assumption def closed (s : set α) : Prop := open' (-s) @[simp] -lemma closed_empty : closed (∅ : set α) := by simp [closed] +theorem closed_empty : closed (∅ : set α) := by simp [closed] @[simp] -lemma closed_univ : closed (univ : set α) := by simp [closed] +theorem closed_univ : closed (univ : set α) := by simp [closed] -lemma closed_union : closed s₁ → closed s₂ → closed (s₁ ∪ s₂) := +theorem closed_union : closed s₁ → closed s₂ → closed (s₁ ∪ s₂) := by simp [closed]; exact open_inter -lemma closed_sInter {s : set (set α)} : (∀t ∈ s, closed t) → closed (⋂₀ s) := +theorem closed_sInter {s : set (set α)} : (∀t ∈ s, closed t) → closed (⋂₀ s) := by simp [closed, compl_sInter]; exact assume h, open_Union $ assume t, open_Union $ assume ht, h t ht -lemma closed_Inter {f : ι → set α} (h : ∀i, closed (f i)) : closed (⋂i, f i ) := +theorem closed_Inter {f : ι → set α} (h : ∀i, closed (f i)) : closed (⋂i, f i ) := closed_sInter $ assume t ⟨i, (heq : t = f i)⟩, heq.symm ▸ h i @[simp] -lemma closed_compl_iff_open {s : set α} : open' (-s) ↔ closed s := +theorem closed_compl_iff_open {s : set α} : open' (-s) ↔ closed s := by refl @[simp] -lemma open_compl_iff_closed {s : set α} : closed (-s) ↔ open' s := +theorem open_compl_iff_closed {s : set α} : closed (-s) ↔ open' s := by rw [←closed_compl_iff_open, compl_compl] -lemma open_diff {s t : set α} (h₁ : open' s) (h₂ : closed t) : open' (s - t) := +theorem open_diff {s t : set α} (h₁ : open' s) (h₂ : closed t) : open' (s - t) := open_inter h₁ $ closed_compl_iff_open.mpr h₂ /- interior -/ def interior (s : set α) : set α := ⋃₀ {t | open' t ∧ t ⊆ s} @[simp] -lemma open_interior {s : set α} : open' (interior s) := +theorem open_interior {s : set α} : open' (interior s) := open_sUnion $ assume t ⟨h₁, h₂⟩, h₁ -lemma interior_subset {s : set α} : interior s ⊆ s := +theorem interior_subset {s : set α} : interior s ⊆ s := sUnion_subset $ assume t ⟨h₁, h₂⟩, h₂ -lemma interior_maximal {s t : set α} (h₁ : t ⊆ s) (h₂ : open' t) : t ⊆ interior s := +theorem interior_maximal {s t : set α} (h₁ : t ⊆ s) (h₂ : open' t) : t ⊆ interior s := subset_sUnion_of_mem ⟨h₂, h₁⟩ -lemma interior_eq_of_open {s : set α} (h : open' s) : interior s = s := +theorem interior_eq_of_open {s : set α} (h : open' s) : interior s = s := subset.antisymm interior_subset (interior_maximal (subset.refl s) h) -lemma interior_eq_iff_open {s : set α} : interior s = s ↔ open' s := +theorem interior_eq_iff_open {s : set α} : interior s = s ↔ open' s := ⟨assume h, h ▸ open_interior, interior_eq_of_open⟩ -lemma subset_interior_iff_subset_of_open {s t : set α} (h₁ : open' s) : +theorem subset_interior_iff_subset_of_open {s t : set α} (h₁ : open' s) : s ⊆ interior t ↔ s ⊆ t := ⟨assume h, subset.trans h interior_subset, assume h₂, interior_maximal h₂ h₁⟩ -lemma interior_mono {s t : set α} (h : s ⊆ t) : interior s ⊆ interior t := +theorem interior_mono {s t : set α} (h : s ⊆ t) : interior s ⊆ interior t := interior_maximal (subset.trans interior_subset h) open_interior @[simp] -lemma interior_empty : interior (∅ : set α) = ∅ := +theorem interior_empty : interior (∅ : set α) = ∅ := interior_eq_of_open open_empty @[simp] -lemma interior_univ : interior (univ : set α) = univ := +theorem interior_univ : interior (univ : set α) = univ := interior_eq_of_open open_univ @[simp] -lemma interior_interior {s : set α} : interior (interior s) = interior s := +theorem interior_interior {s : set α} : interior (interior s) = interior s := interior_eq_of_open open_interior @[simp] -lemma interior_inter {s t : set α} : interior (s ∩ t) = interior s ∩ interior t := +theorem interior_inter {s t : set α} : interior (s ∩ t) = interior s ∩ interior t := subset.antisymm (subset_inter (interior_mono $ inter_subset_left s t) (interior_mono $ inter_subset_right s t)) (interior_maximal (inter_subset_inter interior_subset interior_subset) $ by simp [open_inter]) -lemma interior_union_closed_of_interior_empty {s t : set α} (h₁ : closed s) (h₂ : interior t = ∅) : +theorem interior_union_closed_of_interior_empty {s t : set α} (h₁ : closed s) (h₂ : interior t = ∅) : interior (s ∪ t) = interior s := have interior (s ∪ t) ⊆ s, from assume x ⟨u, ⟨(hu₁ : open' u), (hu₂ : u ⊆ s ∪ t)⟩, (hx₁ : x ∈ u)⟩, @@ -149,50 +149,50 @@ subset.antisymm def closure (s : set α) : set α := ⋂₀ {t | closed t ∧ s ⊆ t} @[simp] -lemma closed_closure {s : set α} : closed (closure s) := +theorem closed_closure {s : set α} : closed (closure s) := closed_sInter $ assume t ⟨h₁, h₂⟩, h₁ -lemma subset_closure {s : set α} : s ⊆ closure s := +theorem subset_closure {s : set α} : s ⊆ closure s := subset_sInter $ assume t ⟨h₁, h₂⟩, h₂ -lemma closure_minimal {s t : set α} (h₁ : s ⊆ t) (h₂ : closed t) : closure s ⊆ t := +theorem closure_minimal {s t : set α} (h₁ : s ⊆ t) (h₂ : closed t) : closure s ⊆ t := sInter_subset_of_mem ⟨h₂, h₁⟩ -lemma closure_eq_of_closed {s : set α} (h : closed s) : closure s = s := +theorem closure_eq_of_closed {s : set α} (h : closed s) : closure s = s := subset.antisymm (closure_minimal (subset.refl s) h) subset_closure -lemma closure_eq_iff_closed {s : set α} : closure s = s ↔ closed s := +theorem closure_eq_iff_closed {s : set α} : closure s = s ↔ closed s := ⟨assume h, h ▸ closed_closure, closure_eq_of_closed⟩ -lemma closure_subset_iff_subset_of_closed {s t : set α} (h₁ : closed t) : +theorem closure_subset_iff_subset_of_closed {s t : set α} (h₁ : closed t) : closure s ⊆ t ↔ s ⊆ t := ⟨subset.trans subset_closure, assume h, closure_minimal h h₁⟩ -lemma closure_mono {s t : set α} (h : s ⊆ t) : closure s ⊆ closure t := +theorem closure_mono {s t : set α} (h : s ⊆ t) : closure s ⊆ closure t := closure_minimal (subset.trans h subset_closure) closed_closure @[simp] -lemma closure_empty : closure (∅ : set α) = ∅ := +theorem closure_empty : closure (∅ : set α) = ∅ := closure_eq_of_closed closed_empty @[simp] -lemma closure_univ : closure (univ : set α) = univ := +theorem closure_univ : closure (univ : set α) = univ := closure_eq_of_closed closed_univ @[simp] -lemma closure_closure {s : set α} : closure (closure s) = closure s := +theorem closure_closure {s : set α} : closure (closure s) = closure s := closure_eq_of_closed closed_closure @[simp] -lemma closure_union {s t : set α} : closure (s ∪ t) = closure s ∪ closure t := +theorem closure_union {s t : set α} : closure (s ∪ t) = closure s ∪ closure t := subset.antisymm (closure_minimal (union_subset_union subset_closure subset_closure) $ by simp [closed_union]) (union_subset (closure_mono $ subset_union_left _ _) (closure_mono $ subset_union_right _ _)) -lemma interior_subset_closure {s : set α} : interior s ⊆ closure s := +theorem interior_subset_closure {s : set α} : interior s ⊆ closure s := subset.trans interior_subset subset_closure -lemma closure_eq_compl_interior_compl {s : set α} : closure s = - interior (- s) := +theorem closure_eq_compl_interior_compl {s : set α} : closure s = - interior (- s) := begin simp [interior, closure], rw [compl_sUnion, compl_image_set_of], @@ -200,17 +200,17 @@ begin end @[simp] -lemma interior_compl_eq {s : set α} : interior (- s) = - closure s := +theorem interior_compl_eq {s : set α} : interior (- s) = - closure s := by simp [closure_eq_compl_interior_compl] @[simp] -lemma closure_compl_eq {s : set α} : closure (- s) = - interior s := +theorem closure_compl_eq {s : set α} : closure (- s) = - interior s := by simp [closure_eq_compl_interior_compl] /- neighbourhood filter -/ def nhds (a : α) : filter α := (⨅ s ∈ {s : set α | a ∈ s ∧ open' s}, principal s) -lemma nhds_sets {a : α} : (nhds a).sets = {s | ∃t⊆s, open' t ∧ a ∈ t} := +theorem nhds_sets {a : α} : (nhds a).sets = {s | ∃t⊆s, open' t ∧ a ∈ t} := calc (nhds a).sets = (⋃s∈{s : set α| a ∈ s ∧ open' s}, (principal s).sets) : infi_sets_eq' begin simp, @@ -223,7 +223,7 @@ calc (nhds a).sets = (⋃s∈{s : set α| a ∈ s ∧ open' s}, (principal s).se (supr_le $ assume i, supr_le $ assume ⟨hi₁, hi₂⟩ t ht, ⟨i, ht, hi₂, hi₁⟩) (assume t ⟨i, hi₁, hi₂, hi₃⟩, begin simp; exact ⟨i, hi₂, hi₁, hi₃⟩ end) -lemma map_nhds {a : α} {f : α → β} : +theorem map_nhds {a : α} {f : α → β} : map f (nhds a) = (⨅ s ∈ {s : set α | a ∈ s ∧ open' s}, principal (image f s)) := calc map f (nhds a) = (⨅ s ∈ {s : set α | a ∈ s ∧ open' s}, map f (principal s)) : map_binfi_eq @@ -235,33 +235,33 @@ calc map f (nhds a) = (⨅ s ∈ {s : set α | a ∈ s ∧ open' s}, map f (prin ⟨univ, by simp⟩ ... = _ : by simp -lemma mem_nhds_sets_iff {a : α} {s : set α} : +theorem mem_nhds_sets_iff {a : α} {s : set α} : s ∈ (nhds a).sets ↔ ∃t⊆s, open' t ∧ a ∈ t := by simp [nhds_sets] -lemma mem_nhds_sets {a : α} {s : set α} (hs : open' s) (ha : a ∈ s) : +theorem mem_nhds_sets {a : α} {s : set α} (hs : open' s) (ha : a ∈ s) : s ∈ (nhds a).sets := by simp [nhds_sets]; exact ⟨s, hs, subset.refl _, ha⟩ -lemma return_le_nhds : return ≤ (nhds : α → filter α) := +theorem return_le_nhds : return ≤ (nhds : α → filter α) := assume a, le_infi $ assume s, le_infi $ assume ⟨h₁, _⟩, principal_mono.mpr $ by simp [h₁] @[simp] -lemma nhds_neq_bot {a : α} : nhds a ≠ ⊥ := +theorem nhds_neq_bot {a : α} : nhds a ≠ ⊥ := assume : nhds a = ⊥, have return a = (⊥ : filter α), from lattice.bot_unique $ this ▸ return_le_nhds a, return_neq_bot this -lemma interior_eq_nhds {s : set α} : interior s = {a | nhds a ≤ principal s} := +theorem interior_eq_nhds {s : set α} : interior s = {a | nhds a ≤ principal s} := set.ext $ by simp [interior, nhds_sets] -lemma open_iff_nhds {s : set α} : open' s ↔ (∀a∈s, nhds a ≤ principal s) := +theorem open_iff_nhds {s : set α} : open' s ↔ (∀a∈s, nhds a ≤ principal s) := calc open' s ↔ interior s = s : by rw [interior_eq_iff_open] ... ↔ s ⊆ interior s : ⟨assume h, by simp [*, subset.refl], subset.antisymm interior_subset⟩ ... ↔ (∀a∈s, nhds a ≤ principal s) : by rw [interior_eq_nhds]; refl -lemma closure_eq_nhds {s : set α} : closure s = {a | nhds a ⊓ principal s ≠ ⊥} := +theorem closure_eq_nhds {s : set α} : closure s = {a | nhds a ⊓ principal s ≠ ⊥} := calc closure s = - interior (- s) : closure_eq_compl_interior_compl ... = {a | ¬ nhds a ≤ principal (-s)} : by rw [interior_eq_nhds]; refl ... = {a | nhds a ⊓ principal s ≠ ⊥} : set.ext $ assume a, not_congr @@ -269,7 +269,7 @@ calc closure s = - interior (- s) : closure_eq_compl_interior_compl (show principal s ⊔ principal (-s) = ⊤, by simp [principal_univ]) (by simp)).symm -lemma closed_iff_nhds {s : set α} : closed s ↔ (∀a, nhds a ⊓ principal s ≠ ⊥ → a ∈ s) := +theorem closed_iff_nhds {s : set α} : closed s ↔ (∀a, nhds a ⊓ principal s ≠ ⊥ → a ∈ s) := calc closed s ↔ closure s = s : by rw [closure_eq_iff_closed] ... ↔ closure s ⊆ s : ⟨assume h, by simp [*, subset.refl], assume h, subset.antisymm h subset_closure⟩ ... ↔ (∀a, nhds a ⊓ principal s ≠ ⊥ → a ∈ s) : by rw [closure_eq_nhds]; refl @@ -284,7 +284,7 @@ theorem not_eq_empty_iff_exists {s : set α} : ¬ (s = ∅) ↔ ∃ x, x ∈ s : ⟨exists_mem_of_ne_empty, assume ⟨x, (hx : x ∈ s)⟩ h_eq, by rw [h_eq] at hx; assumption⟩ -lemma closed_Union_of_locally_finite {f : β → set α} +theorem closed_Union_of_locally_finite {f : β → set α} (h₁ : locally_finite f) (h₂ : ∀i, closed (f i)) : closed (⋃i, f i) := open_iff_nhds.mpr $ assume a, assume h : a ∉ (⋃i, f i), have ∀i, a ∈ -f i, @@ -314,7 +314,7 @@ section compact def compact (s : set α) := ∀{f}, f ≠ ⊥ → f ≤ principal s → ∃a∈s, f ⊓ nhds a ≠ ⊥ -lemma compact_adherence_nhdset {s t : set α} {f : filter α} +theorem compact_adherence_nhdset {s t : set α} {f : filter α} (hs : compact s) (hf₂ : f ≤ principal s) (ht₁ : open' t) (ht₂ : ∀a∈s, nhds a ⊓ f ≠ ⊥ → a ∈ t) : t ∈ f.sets := classical.by_cases mem_sets_of_neq_bot $ @@ -330,7 +330,7 @@ classical.by_cases mem_sets_of_neq_bot $ from this _ ⟨t, mem_nhds_sets ht₁ ‹a ∈ t ›, -t, subset.refl _, subset.refl _⟩ (by simp), by contradiction -lemma compact_iff_ultrafilter_le_nhds {s : set α} : +theorem compact_iff_ultrafilter_le_nhds {s : set α} : compact s ↔ (∀f, ultrafilter f → f ≤ principal s → ∃a∈s, f ≤ nhds a) := ⟨assume hs : compact s, assume f hf hfs, let ⟨a, ha, h⟩ := hs hf.left hfs in @@ -344,7 +344,7 @@ lemma compact_iff_ultrafilter_le_nhds {s : set α} : by simp [inf_of_le_left, h]; exact (ultrafilter_ultrafilter_of hf).left, ⟨a, ha, neq_bot_of_le_neq_bot this (inf_le_inf ultrafilter_of_le (le_refl _))⟩⟩ -lemma finite_subcover_of_compact {s : set α} {c : set (set α)} +theorem finite_subcover_of_compact {s : set α} {c : set (set α)} (hs : compact s) (hc₁ : ∀t∈c, open' t) (hc₂ : s ⊆ ⋃₀ c) : ∃c'⊆c, finite c' ∧ s ⊆ ⋃₀ c' := classical.by_contradiction $ assume h, have h : ∀{c'}, c' ⊆ c → finite c' → ¬ s ⊆ ⋃₀ c', @@ -382,7 +382,7 @@ class t1_space (α : Type u) [topological_space α] := class t2_space (α : Type u) [topological_space α] := (t2 : ∀x y, x ≠ y → ∃u v : set α, open' u ∧ open' v ∧ x ∈ u ∧ y ∈ v ∧ u ∩ v = ∅) -lemma eq_of_nhds_neq_bot [ht : t2_space α] {x y : α} (h : nhds x ⊓ nhds y ≠ ⊥) : x = y := +theorem eq_of_nhds_neq_bot [ht : t2_space α] {x y : α} (h : nhds x ⊓ nhds y ≠ ⊥) : x = y := classical.by_contradiction $ assume : x ≠ y, let ⟨u, v, hu, hv, hx, hy, huv⟩ := t2_space.t2 _ x y this in have h₁ : u ∈ (nhds x ⊓ nhds y).sets, @@ -413,7 +413,7 @@ def generate_from (g : set (set α)) : topological_space α := open_inter := generate_open.inter, open_sUnion := generate_open.sUnion } -lemma nhds_generate_from {g : set (set α)} {a : α} : +theorem nhds_generate_from {g : set (set α)} {a : α} : @nhds α (generate_from g) a = (⨅s∈{s | a ∈ s ∧ s ∈ g}, principal s) := le_antisymm (infi_le_infi $ assume s, infi_le_infi_const $ assume ⟨as, sg⟩, ⟨as, generate_open.basic _ sg⟩) @@ -532,16 +532,16 @@ instance {α : Type u} : complete_lattice (topological_space α) := instance inhabited_topological_space {α : Type u} : inhabited (topological_space α) := ⟨⊤⟩ -lemma t2_space_top : @t2_space α ⊤ := +theorem t2_space_top : @t2_space α ⊤ := ⟨assume x y hxy, ⟨{x}, {y}, trivial, trivial, mem_insert _ _, mem_insert _ _, eq_empty_of_forall_not_mem $ by intros z hz; simp at hz; cc⟩⟩ -lemma le_of_nhds_le_nhds {t₁ t₂ : topological_space α} (h : ∀x, @nhds α t₂ x ≤ @nhds α t₁ x) : +theorem le_of_nhds_le_nhds {t₁ t₂ : topological_space α} (h : ∀x, @nhds α t₂ x ≤ @nhds α t₁ x) : t₁ ≤ t₂ := assume s, show @open' α t₁ s → @open' α t₂ s, begin simp [open_iff_nhds]; exact assume hs a ha, h _ $ hs _ ha end -lemma eq_of_nhds_eq_nhds {t₁ t₂ : topological_space α} (h : ∀x, @nhds α t₂ x = @nhds α t₁ x) : +theorem eq_of_nhds_eq_nhds {t₁ t₂ : topological_space α} (h : ∀x, @nhds α t₂ x = @nhds α t₁ x) : t₁ = t₂ := le_antisymm (le_of_nhds_le_nhds $ assume x, le_of_eq $ h x) @@ -574,14 +574,14 @@ instance topological_space_Pi {β : α → Type v} [t₂ : Πa, topological_spac section open topological_space -lemma generate_from_le {t : topological_space α} { g : set (set α) } (h : ∀s∈g, open' s) : +theorem generate_from_le {t : topological_space α} { g : set (set α) } (h : ∀s∈g, open' s) : generate_from g ≤ t := assume s (hs : generate_open g s), generate_open.rec_on hs h open_univ (assume s t _ _ hs ht, open_inter hs ht) (assume k _ hk, open_sUnion hk) -lemma supr_eq_generate_from {ι : Sort w} { g : ι → topological_space α } : +theorem supr_eq_generate_from {ι : Sort w} { g : ι → topological_space α } : supr g = generate_from (⋃i, {s | (g i).open' s}) := le_antisymm (supr_le $ assume i s open_s, @@ -594,7 +594,7 @@ le_antisymm this s open_s end) -lemma sup_eq_generate_from { g₁ g₂ : topological_space α } : +theorem sup_eq_generate_from { g₁ g₂ : topological_space α } : g₁ ⊔ g₂ = generate_from {s | g₁.open' s ∨ g₂.open' s} := le_antisymm (sup_le (assume s, generate_open.basic _ ∘ or.inl) (assume s, generate_open.basic _ ∘ or.inr)) @@ -603,10 +603,10 @@ le_antisymm have h₂ : g₂ ≤ g₁ ⊔ g₂, from le_sup_right, or.rec_on hs (h₁ s) (h₂ s)) -lemma nhds_mono {t₁ t₂ : topological_space α} {a : α} (h : t₁ ≤ t₂) : @nhds α t₂ a ≤ @nhds α t₁ a := +theorem nhds_mono {t₁ t₂ : topological_space α} {a : α} (h : t₁ ≤ t₂) : @nhds α t₂ a ≤ @nhds α t₁ a := infi_le_infi $ assume s, infi_le_infi2 $ assume ⟨ha, hs⟩, ⟨⟨ha, h _ hs⟩, le_refl _⟩ -lemma nhds_supr {ι : Sort w} {t : ι → topological_space α} {a : α} : +theorem nhds_supr {ι : Sort w} {t : ι → topological_space α} {a : α} : @nhds α (supr t) a = (⨅i, @nhds α (t i) a) := le_antisymm (le_infi $ assume i, nhds_mono $ le_supr _ _) diff --git a/topology/uniform_space.lean b/topology/uniform_space.lean index 8d3becee60ac8..1636f3e07f4ee 100644 --- a/topology/uniform_space.lean +++ b/topology/uniform_space.lean @@ -23,18 +23,18 @@ def id_rel {α : Type u} := {p : α × α | p.1 = p.2} def comp_rel {α : Type u} (r₁ r₂ : set (α×α)) := {p : α × α | ∃z:α, (p.1, z) ∈ r₁ ∧ (z, p.2) ∈ r₂} -@[simp] lemma swap_id_rel : prod.swap '' id_rel = @id_rel α := +@[simp] theorem swap_id_rel : prod.swap '' id_rel = @id_rel α := set.ext $ assume ⟨a, b⟩, by simp [image_swap_eq_vimage_swap]; exact eq_comm -lemma monotone_comp_rel [weak_order β] {f g : β → set (α×α)} +theorem monotone_comp_rel [weak_order β] {f g : β → set (α×α)} (hf : monotone f) (hg : monotone g) : monotone (λx, comp_rel (f x) (g x)) := assume a b h p ⟨z, h₁, h₂⟩, ⟨z, hf h h₁, hg h h₂⟩ -lemma prod_mk_mem_comp_rel {a b c : α} {s t : set (α×α)} (h₁ : (a, c) ∈ s) (h₂ : (c, b) ∈ t) : +theorem prod_mk_mem_comp_rel {a b c : α} {s t : set (α×α)} (h₁ : (a, c) ∈ s) (h₂ : (c, b) ∈ t) : (a, b) ∈ comp_rel s t := ⟨c, h₁, h₂⟩ -@[simp] lemma id_comp_rel {r : set (α×α)} : comp_rel id_rel r = r := +@[simp] theorem id_comp_rel {r : set (α×α)} : comp_rel id_rel r = r := set.ext $ assume ⟨a, b⟩, ⟨assume ⟨a', (heq : a = a'), ha'⟩, heq.symm ▸ ha', assume ha, ⟨a, rfl, ha⟩⟩ /- uniformity -/ @@ -45,7 +45,7 @@ class uniform_space (α : Type u) := (symm : prod.swap <$> uniformity ≤ uniformity) (comp : uniformity.lift' (λs, comp_rel s s) ≤ uniformity) -lemma uniform_space_eq {u₁ u₂ : uniform_space α} (h : u₁.uniformity = u₂.uniformity) : u₁ = u₂ := +theorem uniform_space_eq {u₁ u₂ : uniform_space α} (h : u₁.uniformity = u₂.uniformity) : u₁ = u₂ := begin cases u₁ with a, cases u₂ with b, have h' : a = b, assumption, @@ -58,38 +58,38 @@ variables [uniform_space α] def uniformity : filter (α × α) := uniform_space.uniformity α -lemma refl_le_uniformity : principal id_rel ≤ @uniformity α _ := +theorem refl_le_uniformity : principal id_rel ≤ @uniformity α _ := uniform_space.refl α -lemma refl_mem_uniformity {x : α} {s : set (α × α)} (h : s ∈ (@uniformity α _).sets) : +theorem refl_mem_uniformity {x : α} {s : set (α × α)} (h : s ∈ (@uniformity α _).sets) : (x, x) ∈ s := refl_le_uniformity h rfl -lemma symm_le_uniformity : map (@prod.swap α α) uniformity ≤ uniformity := +theorem symm_le_uniformity : map (@prod.swap α α) uniformity ≤ uniformity := uniform_space.symm α -lemma comp_le_uniformity : +theorem comp_le_uniformity : uniformity.lift' (λs:set (α×α), comp_rel s s) ≤ uniformity := uniform_space.comp α -lemma comp_mem_uniformity_sets {s : set (α × α)} (hs : s ∈ (@uniformity α _).sets) : +theorem comp_mem_uniformity_sets {s : set (α × α)} (hs : s ∈ (@uniformity α _).sets) : ∃t∈(@uniformity α _).sets, comp_rel t t ⊆ s := have s ∈ (uniformity.lift' (λt:set (α×α), comp_rel t t)).sets, from comp_le_uniformity hs, (mem_lift'_iff $ monotone_comp_rel monotone_id monotone_id).mp this -lemma symm_of_uniformity {s : set (α × α)} (hs : s ∈ (@uniformity α _).sets) : +theorem symm_of_uniformity {s : set (α × α)} (hs : s ∈ (@uniformity α _).sets) : ∃t∈(@uniformity α _).sets, (∀a b, (a, b) ∈ t → (b, a) ∈ t) ∧ t ⊆ s := have vimage prod.swap s ∈ (@uniformity α _).sets, from symm_le_uniformity hs, ⟨s ∩ vimage prod.swap s, inter_mem_sets hs this, assume a b ⟨h₁, h₂⟩, ⟨h₂, h₁⟩, inter_subset_left _ _⟩ -lemma comp_symm_of_uniformity {s : set (α × α)} (hs : s ∈ (@uniformity α _).sets) : +theorem comp_symm_of_uniformity {s : set (α × α)} (hs : s ∈ (@uniformity α _).sets) : ∃t∈(@uniformity α _).sets, (∀{a b}, (a, b) ∈ t → (b, a) ∈ t) ∧ comp_rel t t ⊆ s := let ⟨t, ht₁, ht₂⟩ := comp_mem_uniformity_sets hs in let ⟨t', ht', ht'₁, ht'₂⟩ := symm_of_uniformity ht₁ in ⟨t', ht', ht'₁, subset.trans (monotone_comp_rel monotone_id monotone_id ht'₂) ht₂⟩ -lemma uniformity_le_symm : uniformity ≤ map (@prod.swap α α) uniformity := +theorem uniformity_le_symm : uniformity ≤ map (@prod.swap α α) uniformity := calc uniformity = id <$> uniformity : (functor.id_map _).symm ... = (prod.swap.{u u} ∘ prod.swap) <$> uniformity : congr_arg (λf : (α×α)→(α×α), f <$> uniformity) (by apply funext; intro x; cases x; refl) @@ -97,16 +97,16 @@ calc uniformity = id <$> uniformity : (functor.id_map _).symm congr map_compose rfl ... ≤ prod.swap.{u u} <$> uniformity : map_mono symm_le_uniformity -lemma uniformity_eq_symm : uniformity = (@prod.swap α α) <$> uniformity := +theorem uniformity_eq_symm : uniformity = (@prod.swap α α) <$> uniformity := le_antisymm uniformity_le_symm symm_le_uniformity -lemma uniformity_lift_le_swap {g : set (α×α) → filter β} {f : filter β} (hg : monotone g) +theorem uniformity_lift_le_swap {g : set (α×α) → filter β} {f : filter β} (hg : monotone g) (h : uniformity.lift (λs, g (vimage prod.swap s)) ≤ f) : uniformity.lift g ≤ f := le_trans (lift_mono uniformity_le_symm (le_refl _)) (by rw [map_lift_eq2 hg, image_swap_eq_vimage_swap]; exact h) -lemma uniformity_lift_le_comp {f : set (α×α) → filter β} (h : monotone f): +theorem uniformity_lift_le_comp {f : set (α×α) → filter β} (h : monotone f): uniformity.lift (λs, f (comp_rel s s)) ≤ uniformity.lift f := calc uniformity.lift (λs, f (comp_rel s s)) = (uniformity.lift' (λs:set (α×α), comp_rel s s)).lift f : @@ -117,7 +117,7 @@ calc uniformity.lift (λs, f (comp_rel s s)) = end ... ≤ uniformity.lift f : lift_mono comp_le_uniformity (le_refl _) -lemma comp_le_uniformity3 : +theorem comp_le_uniformity3 : uniformity.lift' (λs:set (α×α), comp_rel s (comp_rel s s)) ≤ uniformity := calc uniformity.lift' (λd, comp_rel d (comp_rel d d)) = uniformity.lift (λs, uniformity.lift' (λt:set(α×α), comp_rel s (comp_rel t t))) : @@ -145,7 +145,7 @@ instance uniform_space.to_topological_space : topological_space α := uniformity.upwards_sets (hs t ts x xt) $ assume p ph h, ⟨t, ts, ph h⟩ } -lemma mem_nhds_uniformity_iff {x : α} {s : set α} : +theorem mem_nhds_uniformity_iff {x : α} {s : set α} : (s ∈ (nhds x).sets) ↔ ({p : α × α | p.1 = x → p.2 ∈ s} ∈ (@uniformity α _).sets) := ⟨ begin simp [mem_nhds_sets_iff], @@ -171,7 +171,7 @@ lemma mem_nhds_uniformity_iff {x : α} {s : set α} : from tr this rfl, hs⟩⟩ -lemma nhds_eq_uniformity {x : α} : +theorem nhds_eq_uniformity {x : α} : nhds x = uniformity.lift' (λs:set (α×α), {y | (x, y) ∈ s}) := filter_eq $ set.ext $ assume s, begin @@ -184,17 +184,17 @@ filter_eq $ set.ext $ assume s, show (x, y) ∈ t, from eq ▸ hp⟩ end -lemma mem_nhds_left {x : α} {s : set (α×α)} (h : s ∈ (uniformity.sets : set (set (α×α)))) : +theorem mem_nhds_left {x : α} {s : set (α×α)} (h : s ∈ (uniformity.sets : set (set (α×α)))) : {y : α | (x, y) ∈ s} ∈ (nhds x).sets := have nhds x ≤ principal {y : α | (x, y) ∈ s}, by rw [nhds_eq_uniformity]; exact infi_le_of_le s (infi_le _ h), by simp at this; assumption -lemma mem_nhds_right {y : α} {s : set (α×α)} (h : s ∈ (uniformity.sets : set (set (α×α)))) : +theorem mem_nhds_right {y : α} {s : set (α×α)} (h : s ∈ (uniformity.sets : set (set (α×α)))) : {x : α | (x, y) ∈ s} ∈ (nhds y).sets := mem_nhds_left (symm_le_uniformity h) -lemma lift_nhds_left {x : α} {g : set α → filter β} (hg : monotone g) : +theorem lift_nhds_left {x : α} {g : set α → filter β} (hg : monotone g) : (nhds x).lift g = uniformity.lift (λs:set (α×α), g {y | (x, y) ∈ s}) := eq.trans begin @@ -203,7 +203,7 @@ eq.trans end (congr_arg _ $ funext $ assume s, filter.lift_principal hg) -lemma lift_nhds_right {x : α} {g : set α → filter β} (hg : monotone g) : +theorem lift_nhds_right {x : α} {g : set α → filter β} (hg : monotone g) : (nhds x).lift g = uniformity.lift (λs:set (α×α), g {y | (y, x) ∈ s}) := calc (nhds x).lift g = uniformity.lift (λs:set (α×α), g {y | (x, y) ∈ s}) : lift_nhds_left hg ... = ((@prod.swap α α) <$> uniformity).lift (λs:set (α×α), g {y | (x, y) ∈ s}) : by rw [←uniformity_eq_symm] @@ -211,7 +211,7 @@ calc (nhds x).lift g = uniformity.lift (λs:set (α×α), g {y | (x, y) ∈ s}) map_lift_eq2 $ monotone_comp monotone_vimage hg ... = _ : by simp [image_swap_eq_vimage_swap] -lemma nhds_nhds_eq_uniformity_uniformity_prod {a b : α} : +theorem nhds_nhds_eq_uniformity_uniformity_prod {a b : α} : filter.prod (nhds a) (nhds b) = uniformity.lift (λs:set (α×α), uniformity.lift' (λt:set (α×α), set.prod {y : α | (y, a) ∈ s} {y : α | (b, y) ∈ t})) := @@ -226,7 +226,7 @@ begin assume x, monotone_prod monotone_id monotone_const) end -lemma nhds_eq_uniformity_prod {a b : α} : +theorem nhds_eq_uniformity_prod {a b : α} : nhds (a, b) = uniformity.lift' (λs:set (α×α), set.prod {y : α | (y, a) ∈ s} {y : α | (b, y) ∈ s}) := begin @@ -235,7 +235,7 @@ begin { intro t, exact monotone_prod monotone_vimage monotone_const } end -lemma nhdset_of_mem_uniformity {d : set (α×α)} (s : set (α×α)) (hd : d ∈ (@uniformity α _).sets) : +theorem nhdset_of_mem_uniformity {d : set (α×α)} (s : set (α×α)) (hd : d ∈ (@uniformity α _).sets) : ∃(t : set (α×α)), open' t ∧ s ⊆ t ∧ t ⊆ {p | ∃x y, (p.1, x) ∈ d ∧ (x, y) ∈ s ∧ (y, p.2) ∈ d} := let cl_d := {p:α×α | ∃x y, (p.1, x) ∈ d ∧ (x, y) ∈ s ∧ (y, p.2) ∈ d} in have ∀p ∈ s, ∃t ⊆ cl_d, open' t ∧ p ∈ t, from @@ -257,7 +257,7 @@ match this with Union_subset $ assume p, Union_subset $ assume hp, (ht p hp).left⟩ end -lemma closure_eq_inter_uniformity {t : set (α×α)} : +theorem closure_eq_inter_uniformity {t : set (α×α)} : closure t = (⋂ d∈(@uniformity α _).sets, comp_rel d (comp_rel t d)) := set.ext $ assume ⟨a, b⟩, calc (a, b) ∈ closure t ↔ (nhds (a, b) ⊓ principal t ≠ ⊥) : by simp [closure_eq_nhds] @@ -286,14 +286,14 @@ calc (a, b) ∈ closure t ↔ (nhds (a, b) ⊓ principal t ≠ ⊥) : by simp [c assume ⟨x, hx, y, hxyt, hy⟩, ⟨⟨x, y⟩, ⟨⟨hx, hy⟩, hxyt⟩⟩⟩ ... ↔ _ : by simp -lemma uniformity_eq_uniformity_closure : (@uniformity α _) = uniformity.lift' closure := +theorem uniformity_eq_uniformity_closure : (@uniformity α _) = uniformity.lift' closure := le_antisymm (le_infi $ assume s, le_infi $ assume hs, by simp; exact uniformity.upwards_sets hs subset_closure) (calc uniformity.lift' closure ≤ uniformity.lift' (λd, comp_rel d (comp_rel d d)) : lift'_mono' (by intros s hs; rw [closure_eq_inter_uniformity]; exact bInter_subset_of_mem hs) ... ≤ uniformity : comp_le_uniformity3) -lemma uniformity_eq_uniformity_interior : (@uniformity α _) = uniformity.lift' interior := +theorem uniformity_eq_uniformity_interior : (@uniformity α _) = uniformity.lift' interior := le_antisymm (le_infi $ assume d, le_infi $ assume hd, let ⟨s, hs, hs_comp⟩ := (mem_lift'_iff $ @@ -308,7 +308,7 @@ le_antisymm by simp [this]) (assume s hs, (uniformity.lift' interior).upwards_sets (mem_lift' hs) interior_subset) -lemma interior_mem_uniformity {s : set (α × α)} (hs : s ∈ (@uniformity α _).sets) : +theorem interior_mem_uniformity {s : set (α × α)} (hs : s ∈ (@uniformity α _).sets) : interior s ∈ (@uniformity α _).sets := by rw [uniformity_eq_uniformity_interior]; exact mem_lift' hs @@ -321,11 +321,11 @@ def uniform_embedding [uniform_space β] (f : α → β) := (∀a₁ a₂, f a₁ = f a₂ → a₁ = a₂) ∧ vmap (λx:α×α, (f x.1, f x.2)) uniformity = uniformity -lemma uniform_continuous_of_embedding [uniform_space β] {f : α → β} +theorem uniform_continuous_of_embedding [uniform_space β] {f : α → β} (hf : uniform_embedding f) : uniform_continuous f := by simp [uniform_continuous, hf.right.symm]; exact assume s hs, ⟨s, hs, subset.refl _⟩ -lemma continuous_of_uniform [uniform_space β] {f : α → β} +theorem continuous_of_uniform [uniform_space β] {f : α → β} (hf : uniform_continuous f) : continuous f := continuous_iff_towards.mpr $ assume a, calc map f (nhds a) ≤ @@ -343,10 +343,10 @@ calc map f (nhds a) ≤ /- cauchy filters -/ def cauchy (f : filter α) := f ≠ ⊥ ∧ filter.prod f f ≤ uniformity -lemma cauchy_downwards {f g : filter α} (h_c : cauchy f) (hg : g ≠ ⊥) (h_le : g ≤ f) : cauchy g := +theorem cauchy_downwards {f g : filter α} (h_c : cauchy f) (hg : g ≠ ⊥) (h_le : g ≤ f) : cauchy g := ⟨hg, le_trans (filter.prod_mono h_le h_le) h_c.right⟩ -lemma cauchy_nhds {a : α} : cauchy (nhds a) := +theorem cauchy_nhds {a : α} : cauchy (nhds a) := ⟨nhds_neq_bot, calc filter.prod (nhds a) (nhds a) = uniformity.lift (λs:set (α×α), uniformity.lift' (λt:set(α×α), @@ -358,12 +358,12 @@ lemma cauchy_nhds {a : α} : cauchy (nhds a) := assume ⟨x, y⟩ ⟨(hx : (x, a) ∈ s), (hy : (a, y) ∈ s)⟩, ⟨a, hx, hy⟩ ... ≤ uniformity : comp_le_uniformity⟩ -lemma cauchy_pure {a : α} : cauchy (pure a) := +theorem cauchy_pure {a : α} : cauchy (pure a) := cauchy_downwards cauchy_nhds (show principal {a} ≠ ⊥, by simp) (return_le_nhds a) -lemma le_nhds_of_cauchy_adhp {f : filter α} {x : α} (hf : cauchy f) +theorem le_nhds_of_cauchy_adhp {f : filter α} {x : α} (hf : cauchy f) (adhs : f ⊓ nhds x ≠ ⊥) : f ≤ nhds x := have ∀s∈f.sets, x ∈ closure s, begin @@ -399,12 +399,12 @@ calc f ≤ f.lift' (λs:set α, {y | x ∈ closure s ∧ y ∈ closure s}) : ... = nhds x : by rw [nhds_eq_uniformity] -lemma le_nhds_iff_adhp_of_cauchy {f : filter α} {x : α} (hf : cauchy f) : +theorem le_nhds_iff_adhp_of_cauchy {f : filter α} {x : α} (hf : cauchy f) : f ≤ nhds x ↔ f ⊓ nhds x ≠ ⊥ := ⟨assume h, (inf_of_le_left h).symm ▸ hf.left, le_nhds_of_cauchy_adhp hf⟩ -lemma cauchy_map [uniform_space β] {f : filter α} {m : α → β} +theorem cauchy_map [uniform_space β] {f : filter α} {m : α → β} (hm : uniform_continuous m) (hf : cauchy f) : cauchy (map m f) := ⟨have f ≠ ⊥, from hf.left, by simp; assumption, calc filter.prod (map m f) (map m f) = @@ -412,7 +412,7 @@ lemma cauchy_map [uniform_space β] {f : filter α} {m : α → β} ... ≤ map (λp:α×α, (m p.1, m p.2)) uniformity : map_mono hf.right ... ≤ uniformity : hm⟩ -lemma cauchy_vmap [uniform_space β] {f : filter β} {m : α → β} +theorem cauchy_vmap [uniform_space β] {f : filter β} {m : α → β} (hm : vmap (λp:α×α, (m p.1, m p.2)) uniformity ≤ uniformity) (hf : cauchy f) (hb : vmap m f ≠ ⊥) : cauchy (vmap m f) := ⟨hb, @@ -426,7 +426,7 @@ lemma cauchy_vmap [uniform_space β] {f : filter β} {m : α → β} protected def separation_rel (α : Type u) [uniform_space α] := (⋂₀ (@uniformity α _).sets) -lemma separated_equiv : equivalence (λx y, (x, y) ∈ separation_rel α) := +theorem separated_equiv : equivalence (λx y, (x, y) ∈ separation_rel α) := ⟨assume x, assume s, refl_mem_uniformity, assume x y, assume h (s : set (α×α)) hs, have vimage prod.swap s ∈ (@uniformity α _).sets, @@ -470,7 +470,7 @@ have u ∩ v = ∅, from def totally_bounded (s : set α) : Prop := ∀d ∈ (@uniformity α _).sets, ∃t : set α, finite t ∧ s ⊆ (⋃y∈t, {x | (x,y) ∈ d}) -lemma cauchy_of_totally_bounded_of_ultrafilter {s : set α} {f : filter α} +theorem cauchy_of_totally_bounded_of_ultrafilter {s : set α} {f : filter α} (hs : totally_bounded s) (hf : ultrafilter f) (h : f ≤ principal s) : cauchy f := ⟨hf.left, assume t ht, let ⟨t', ht'₁, ht'_symm, ht'_t⟩ := comp_symm_of_uniformity ht in @@ -485,7 +485,7 @@ lemma cauchy_of_totally_bounded_of_ultrafilter {s : set α} {f : filter α} ⟨y, h₁, ht'_symm h₂⟩, (filter.prod f f).upwards_sets (prod_mem_prod hif hif) (subset.trans this ht'_t)⟩ -lemma totally_bounded_iff_filter {s : set α} : +theorem totally_bounded_iff_filter {s : set α} : totally_bounded s ↔ (∀f, f ≠ ⊥ → f ≤ principal s → ∃c ≤ f, cauchy c) := ⟨assume : totally_bounded s, assume f hf hs, ⟨ultrafilter_of f, ultrafilter_of_le, @@ -531,7 +531,7 @@ lemma totally_bounded_iff_filter {s : set α} : from c.upwards_sets this $ assume x ⟨⟨hxs, hxys⟩, hxm, _⟩, hxys $ ‹m ⊆ ys› hxm, hc₂.left $ empty_in_sets_eq_bot.mp this⟩ -lemma totally_bounded_iff_ultrafilter {s : set α} : +theorem totally_bounded_iff_ultrafilter {s : set α} : totally_bounded s ↔ (∀f, ultrafilter f → f ≤ principal s → cauchy f) := ⟨assume hs f, cauchy_of_totally_bounded_of_ultrafilter hs, assume h, totally_bounded_iff_filter.mpr $ assume f hf hfs, @@ -539,7 +539,7 @@ lemma totally_bounded_iff_ultrafilter {s : set α} : from h (ultrafilter_of f) (ultrafilter_ultrafilter_of hf) (le_trans ultrafilter_of_le hfs), ⟨ultrafilter_of f, ultrafilter_of_le, this⟩⟩ -lemma compact_of_totally_bounded_complete {s : set α} +theorem compact_of_totally_bounded_complete {s : set α} (ht : totally_bounded s) (hc : ∀{f:filter α}, cauchy f → f ≤ principal s → ∃x∈s, f ≤ nhds x) : compact s := begin @@ -552,18 +552,18 @@ end class complete_space (α : Type u) [uniform_space α] : Prop := (complete : ∀{f:filter α}, cauchy f → ∃x, f ≤ nhds x) -lemma complete_of_closed [complete_space α] {s : set α} {f : filter α} +theorem complete_of_closed [complete_space α] {s : set α} {f : filter α} (h : closed s) (hf : cauchy f) (hfs : f ≤ principal s) : ∃x∈s, f ≤ nhds x := let ⟨x, hx⟩ := complete_space.complete hf in have x ∈ s, from closed_iff_nhds.mp h x $ neq_bot_of_le_neq_bot hf.left $ le_inf hx hfs, ⟨x, this, hx⟩ -lemma compact_of_totally_bounded_closed [complete_space α] {s : set α} +theorem compact_of_totally_bounded_closed [complete_space α] {s : set α} (ht : totally_bounded s) (hc : closed s) : compact s := @compact_of_totally_bounded_complete α _ s ht $ assume f, complete_of_closed hc -lemma complete_space_extension [uniform_space β] {m : β → α} +theorem complete_space_extension [uniform_space β] {m : β → α} (hm : uniform_embedding m) (dense : ∀x, x ∈ closure (m '' univ)) (h : ∀f:filter β, cauchy f → ∃x:α, map m f ≤ nhds x) : @@ -669,11 +669,11 @@ instance : uniform_space (quotient (separation_setoid α)) := ... ≤ map (λp:(α×α), (⟦p.1⟧, ⟦p.2⟧)) uniformity : map_mono comp_le_uniformity3 } -lemma uniform_continuous_quotient_mk : +theorem uniform_continuous_quotient_mk : uniform_continuous (quotient.mk : α → quotient (separation_setoid α)) := le_refl _ -lemma vmap_quotient_le_uniformity : vmap (λ (p : α × α), (⟦p.fst⟧, ⟦p.snd⟧)) uniformity ≤ uniformity := +theorem vmap_quotient_le_uniformity : vmap (λ (p : α × α), (⟦p.fst⟧, ⟦p.snd⟧)) uniformity ≤ uniformity := assume t' ht', let ⟨t, ht, tt_t'⟩ := comp_mem_uniformity_sets ht' in let ⟨s, hs, ss_t⟩ := comp_mem_uniformity_sets ht in @@ -687,7 +687,7 @@ let ⟨s, hs, ss_t⟩ := comp_mem_uniformity_sets ht in tt_t' ⟨b₁, show ((a₁, a₂).1, b₁) ∈ t, from ab₁, ss_t ⟨b₂, show ((b₁, a₂).1, b₂) ∈ s, from hb, ba₂⟩⟩⟩ -lemma complete_space_separation [h : complete_space α] : +theorem complete_space_separation [h : complete_space α] : complete_space (quotient (separation_setoid α)) := ⟨assume f, assume hf : cauchy f, have cauchy (vmap (λx, ⟦x⟧) f), from @@ -735,7 +735,7 @@ ne_empty_of_mem $ hs $ show e y ∈ t, from y_eq.symm ▸ hx₁ include h_f -lemma uniformly_extend_spec [nonempty γ] [complete_space γ] {a : α} : +theorem uniformly_extend_spec [nonempty γ] [complete_space γ] {a : α} : map f (vmap e (nhds a)) ≤ nhds (ψ a) := have cauchy (nhds a), from cauchy_nhds, have cauchy (vmap e (nhds a)), from @@ -746,7 +746,7 @@ have ∃c, map f (vmap e (nhds a)) ≤ nhds c, from complete_space.complete this, classical.epsilon_spec this -lemma uniformly_extend_unique [nonempty γ] [cγ : complete_space γ] [sγ : separated γ] {a : α} {c : γ} +theorem uniformly_extend_unique [nonempty γ] [cγ : complete_space γ] [sγ : separated γ] {a : α} {c : γ} (h : map f (vmap e (nhds a)) ≤ nhds c) : ψ a = c := have map f (vmap e (nhds a)) ≤ nhds (ψ a) ⊓ nhds c, from le_inf (@uniformly_extend_spec α β γ _ _ _ _ h_e h_dense f h_f _ cγ a) h, @@ -755,7 +755,7 @@ have nhds (uniformly_extend e f a) ⊓ nhds c ≠ ⊥, from neq_bot_of_le_neq_bot (by simp [map_eq_bot_iff]; exact vmap_e_neq_empty h_e h_dense) this, @eq_of_nhds_neq_bot _ _ (@separated_t2 _ _ sγ) _ _ this -lemma uniformly_extend_of_emb [nonempty γ] [cγ : complete_space γ] [sγ : separated γ] {b : β} : +theorem uniformly_extend_of_emb [nonempty γ] [cγ : complete_space γ] [sγ : separated γ] {b : β} : ψ (e b) = f b := @uniformly_extend_unique α β γ _ _ _ e h_e h_dense f h_f _ cγ sγ (e b) (f b) $ have vmap e (nhds (e b)) ≤ nhds b, @@ -772,7 +772,7 @@ have vmap e (nhds (e b)) ≤ nhds b, calc map f (vmap e (nhds (e b))) ≤ map f (nhds b) : map_mono this ... ≤ nhds (f b) : continuous_iff_towards.mp (continuous_of_uniform h_f) b -lemma uniform_continuous_uniformly_extend [nonempty γ] [cγ : complete_space γ] [sγ : separated γ] : +theorem uniform_continuous_uniformly_extend [nonempty γ] [cγ : complete_space γ] [sγ : separated γ] : uniform_continuous ψ := assume d hd, let ⟨s, hs, (hs_comp : comp_rel s (comp_rel s s) ⊆ d)⟩ := (mem_lift'_iff $ @@ -834,7 +834,7 @@ parameters {α : Type u} [uniform_space α] def gen (s : set (α × α)) : set (Cauchy α × Cauchy α) := {p | s ∈ (filter.prod (p.1.val) (p.2.val)).sets } -lemma monotone_gen : monotone gen := +theorem monotone_gen : monotone gen := monotone_set_of $ assume p, @monotone_mem_sets (α×α) (filter.prod (p.1.val) (p.2.val)) private lemma symm_gen : map prod.swap (uniformity.lift' gen) ≤ uniformity.lift' gen := @@ -902,7 +902,7 @@ instance completion_space : uniform_space (Cauchy α) := def pure_cauchy (a : α) : Cauchy α := ⟨pure a, cauchy_pure⟩ -lemma uniform_embedding_pure_cauchy : uniform_embedding (pure_cauchy : α → Cauchy α) := +theorem uniform_embedding_pure_cauchy : uniform_embedding (pure_cauchy : α → Cauchy α) := ⟨assume a₁ a₂ h, have (pure_cauchy a₁).val = (pure_cauchy a₂).val, from congr_arg _ h, have {a₁} = ({a₂} : set α), @@ -917,7 +917,7 @@ lemma uniform_embedding_pure_cauchy : uniform_embedding (pure_cauchy : α → Ca vmap_lift'_eq monotone_gen ... = uniformity : by simp [this]⟩ -lemma pure_cauchy_dense : ∀x, x ∈ closure (pure_cauchy '' univ) := +theorem pure_cauchy_dense : ∀x, x ∈ closure (pure_cauchy '' univ) := assume f, have h_ex : ∀s∈(@uniformity (Cauchy α) _).sets, ∃y:α, (f, pure_cauchy y) ∈ s, from assume s hs, @@ -1043,11 +1043,11 @@ def uniform_space.vmap (f : α → β) (u : uniform_space β) : uniform_space α end (vmap_mono u.comp) } -lemma uniform_continuous_vmap {f : α → β} {u : uniform_space β} : +theorem uniform_continuous_vmap {f : α → β} {u : uniform_space β} : @uniform_continuous α β (uniform_space.vmap f u) u f := map_vmap_le -lemma to_topological_space_vmap {f : α → β} {u : uniform_space β} : +theorem to_topological_space_vmap {f : α → β} {u : uniform_space β} : @uniform_space.to_topological_space _ (uniform_space.vmap f u) = topological_space.induced f (@uniform_space.to_topological_space β u) := eq_of_nhds_eq_nhds $ assume a, @@ -1074,23 +1074,23 @@ uniform_space.vmap prod.fst t₁ ⊔ uniform_space.vmap prod.snd t₂ /- a similar product space is possible on the function space (uniformity of pointwise convergence), but we want to have the uniformity of uniform convergence on function spaces -/ -lemma to_topological_space_mono {u₁ u₂ : uniform_space α} (h : u₁ ≤ u₂) : +theorem to_topological_space_mono {u₁ u₂ : uniform_space α} (h : u₁ ≤ u₂) : @uniform_space.to_topological_space _ u₁ ≤ @uniform_space.to_topological_space _ u₂ := le_of_nhds_le_nhds $ assume a, by rw [@nhds_eq_uniformity α u₁ a, @nhds_eq_uniformity α u₂ a]; exact (lift'_mono h $ le_refl _) -lemma supr_uniformity {ι : Sort v} {u : ι → uniform_space α} : +theorem supr_uniformity {ι : Sort v} {u : ι → uniform_space α} : (supr u).uniformity = (⨅i, (u i).uniformity) := show (⨅a (h : ∃i:ι, a = u i), a.uniformity) = _, from le_antisymm (le_infi $ assume i, infi_le_of_le (u i) $ infi_le _ ⟨i, rfl⟩) (le_infi $ assume a, le_infi $ assume ⟨i, (ha : a = u i)⟩, ha.symm ▸ infi_le _ _) -lemma to_topological_space_top : @uniform_space.to_topological_space α ⊤ = ⊤ := +theorem to_topological_space_top : @uniform_space.to_topological_space α ⊤ = ⊤ := top_unique $ assume s hs x hx ⟨a₁, a₂⟩ (h₁ : a₁ = a₂) (h₂ : a₁ = x), h₁ ▸ h₂.symm ▸ hx -lemma to_topological_space_bot : @uniform_space.to_topological_space α ⊥ = ⊥ := +theorem to_topological_space_bot : @uniform_space.to_topological_space α ⊥ = ⊥ := bot_unique $ assume s hs, classical.by_cases (assume : s = ∅, this.symm ▸ @open_empty _ ⊥) (assume : s ≠ ∅, @@ -1101,7 +1101,7 @@ bot_unique $ assume s hs, classical.by_cases from top_unique $ assume y hy, @this (x, y) ⟨⟩ rfl, this.symm ▸ @open_univ _ ⊥) -lemma to_topological_space_supr {ι : Sort v} {u : ι → uniform_space α} : +theorem to_topological_space_supr {ι : Sort v} {u : ι → uniform_space α} : @uniform_space.to_topological_space α (supr u) = (⨆i, @uniform_space.to_topological_space α (u i)) := classical.by_cases (assume h : nonempty ι,