From 39129f112b5094bf4750958bead994b4ac9b9e50 Mon Sep 17 00:00:00 2001 From: Jeremy Avigad Date: Mon, 25 May 2015 22:13:23 +1000 Subject: [PATCH] refactor(library/*): do various renamings --- library/algebra/algebra.md | 1 + library/algebra/category/constructions.lean | 12 ++++++------ library/data/data.md | 1 + library/data/default.lean | 2 +- library/data/nat/div.lean | 3 ++- library/data/nat/order.lean | 3 ++- library/data/prod.lean | 2 +- library/data/quotient/basic.lean | 2 +- library/data/quotient/util.lean | 4 ++-- library/data/sigma.lean | 2 +- library/data/subtype.lean | 2 +- library/data/unit.lean | 8 ++++---- library/init/quot.lean | 2 +- library/init/sigma.lean | 2 +- 14 files changed, 25 insertions(+), 21 deletions(-) diff --git a/library/algebra/algebra.md b/library/algebra/algebra.md index 7b53237b2..7fb0861b7 100644 --- a/library/algebra/algebra.md +++ b/library/algebra/algebra.md @@ -14,5 +14,6 @@ Algebraic structures. * [ordered_group](ordered_group.lean) * [ordered_ring](ordered_ring.lean) * [field](field.lean) +* [ordered_field](ordered_field.lean) * [category](category/category.md) : category theory diff --git a/library/algebra/category/constructions.lean b/library/algebra/category/constructions.lean index aad183b3b..474c15d43 100644 --- a/library/algebra/category/constructions.lean +++ b/library/algebra/category/constructions.lean @@ -111,9 +111,9 @@ namespace category mk (λa b, hom (pr1 a) (pr1 b) × hom (pr2 a) (pr2 b)) (λ a b c g f, (pr1 g ∘ pr1 f , pr2 g ∘ pr2 f) ) (λ a, (id,id)) - (λ a b c d h g f, pair_eq !assoc !assoc ) - (λ a b f, prod.equal !id_left !id_left ) - (λ a b f, prod.equal !id_right !id_right) + (λ a b c d h g f, pair_eq !assoc !assoc ) + (λ a b f, prod.eq !id_left !id_left ) + (λ a b f, prod.eq !id_right !id_right) definition Prod_category [reducible] (C D : Category) : Category := Mk (prod_category C D) end @@ -204,9 +204,9 @@ namespace category ... = slice.ob_hom a : {slice.commute f} qed)) (λ a, sigma.mk id !id_right) - (λ a b c d h g f, dpair_eq !assoc !proof_irrel) - (λ a b f, sigma.equal !id_left !proof_irrel) - (λ a b f, sigma.equal !id_right !proof_irrel) + (λ a b c d h g f, dpair_eq !assoc !proof_irrel) + (λ a b f, sigma.eq !id_left !proof_irrel) + (λ a b f, sigma.eq !id_right !proof_irrel) -- We use !proof_irrel instead of rfl, to give the unifier an easier time -- definition slice_category {ob : Type} (C : category ob) (c : ob) : category (Σ(b : ob), hom b c) diff --git a/library/data/data.md b/library/data/data.md index faad28f44..e9e37c8da 100644 --- a/library/data/data.md +++ b/library/data/data.md @@ -27,6 +27,7 @@ Constructors: * [squash](squash.lean) : propositional truncation * [list](list/list.md) * [finset](finset/finset.md) : finite sets +* [stream](stream.lean) * [set](set/set.md) * [vector](vector.lean) diff --git a/library/data/default.lean b/library/data/default.lean index 07b6ba271..77b3e597e 100644 --- a/library/data/default.lean +++ b/library/data/default.lean @@ -4,4 +4,4 @@ Released under Apache 2.0 license as described in the file LICENSE. Author: Jeremy Avigad -/ import .empty .unit .bool .num .string .nat .int .rat .fintype -import .prod .sum .sigma .option .subtype .quotient .list .vector .finset .set +import .prod .sum .sigma .option .subtype .quotient .list .vector .finset .set .stream diff --git a/library/data/nat/div.lean b/library/data/nat/div.lean index 430cacbec..04a5602f3 100644 --- a/library/data/nat/div.lean +++ b/library/data/nat/div.lean @@ -613,7 +613,8 @@ or.elim (eq_zero_or_pos m) theorem eq_zero_of_gcd_eq_zero_right {m n : ℕ} (H : gcd m n = 0) : n = 0 := eq_zero_of_gcd_eq_zero_left (!gcd.comm ▸ H) -theorem gcd_div {m n k : ℕ} (H1 : (k ∣ m)) (H2 : (k ∣ n)) : gcd (m div k) (n div k) = gcd m n div k := +theorem gcd_div {m n k : ℕ} (H1 : (k ∣ m)) (H2 : (k ∣ n)) : + gcd (m div k) (n div k) = gcd m n div k := or.elim (eq_zero_or_pos k) (assume H3 : k = 0, calc diff --git a/library/data/nat/order.lean b/library/data/nat/order.lean index 58fad9c91..c805490b6 100644 --- a/library/data/nat/order.lean +++ b/library/data/nat/order.lean @@ -342,7 +342,8 @@ nat.strong_induction_on a /- pos -/ -theorem by_cases_zero_pos {P : ℕ → Prop} (y : ℕ) (H0 : P 0) (H1 : ∀ {y : nat}, y > 0 → P y) : P y := +theorem by_cases_zero_pos {P : ℕ → Prop} (y : ℕ) (H0 : P 0) (H1 : ∀ {y : nat}, y > 0 → P y) : + P y := nat.cases_on y H0 (take y, H1 !succ_pos) theorem eq_zero_or_pos (n : ℕ) : n = 0 ∨ n > 0 := diff --git a/library/data/prod.lean b/library/data/prod.lean index a1369da43..36c6076dc 100644 --- a/library/data/prod.lean +++ b/library/data/prod.lean @@ -12,7 +12,7 @@ namespace prod theorem pair_eq : a₁ = a₂ → b₁ = b₂ → (a₁, b₁) = (a₂, b₂) := assume H1 H2, H1 ▸ H2 ▸ rfl - protected theorem equal {p₁ p₂ : prod A B} : pr₁ p₁ = pr₁ p₂ → pr₂ p₁ = pr₂ p₂ → p₁ = p₂ := + protected theorem eq {p₁ p₂ : prod A B} : pr₁ p₁ = pr₁ p₂ → pr₂ p₁ = pr₂ p₂ → p₁ = p₂ := destruct p₁ (take a₁ b₁, destruct p₂ (take a₂ b₂ H₁ H₂, pair_eq H₁ H₂)) protected definition is_inhabited [instance] [h₁ : inhabited A] [h₂ : inhabited B] : inhabited (prod A B) := diff --git a/library/data/quotient/basic.lean b/library/data/quotient/basic.lean index 872e1429e..1cacf6da7 100644 --- a/library/data/quotient/basic.lean +++ b/library/data/quotient/basic.lean @@ -284,7 +284,7 @@ intro ... = f (f a) : {Ha⁻¹} ... = f a : representative_map_idempotent H1 H2 a ... = elt_of u : Ha, - show abs (elt_of u) = u, from subtype.equal H) + show abs (elt_of u) = u, from subtype.eq H) (take u : image f, show R (elt_of u) (elt_of u), from obtain (a : A) (Ha : f a = elt_of u), from image_elt_of u, diff --git a/library/data/quotient/util.lean b/library/data/quotient/util.lean index 44722ec5f..fa36ae7c9 100644 --- a/library/data/quotient/util.lean +++ b/library/data/quotient/util.lean @@ -148,7 +148,7 @@ have Hy : pr2 (map_pair2 f v (pair e e)) = pr2 v, from pr2 (map_pair2 f v (pair e e)) = f (pr2 v) (pr2 (pair e e)) : by simp ... = f (pr2 v) e : by simp ... = pr2 v : Hid (pr2 v)), -prod.equal Hx Hy +prod.eq Hx Hy theorem map_pair2_id_left {A B : Type} {f : B → A → A} {e : B} (Hid : ∀a : A, f e a = a) (v : A × A) : map_pair2 f (pair e e) v = v := @@ -162,6 +162,6 @@ have Hy : pr2 (map_pair2 f (pair e e) v) = pr2 v, from pr2 (map_pair2 f (pair e e) v) = f (pr2 (pair e e)) (pr2 v) : by simp ... = f e (pr2 v) : by simp ... = pr2 v : Hid (pr2 v), -prod.equal Hx Hy +prod.eq Hx Hy end quotient diff --git a/library/data/sigma.lean b/library/data/sigma.lean index 5658e1e54..4cf25c0f9 100644 --- a/library/data/sigma.lean +++ b/library/data/sigma.lean @@ -19,7 +19,7 @@ namespace sigma (HB : B == B') (Ha : a == a') (Hb : b == b') : ⟨a, b⟩ == ⟨a', b'⟩ := hcongr_arg4 @mk (heq.type_eq Ha) HB Ha Hb - protected theorem hequal {p : Σa : A, B a} {p' : Σa' : A', B' a'} (HB : B == B') : + protected theorem heq {p : Σa : A, B a} {p' : Σa' : A', B' a'} (HB : B == B') : ∀(H₁ : p.1 == p'.1) (H₂ : p.2 == p'.2), p == p' := destruct p (take a₁ b₁, destruct p' (take a₂ b₂ H₁ H₂, dpair_heq HB H₁ H₂)) diff --git a/library/data/subtype.lean b/library/data/subtype.lean index 39508d099..ae7e5229d 100644 --- a/library/data/subtype.lean +++ b/library/data/subtype.lean @@ -22,7 +22,7 @@ namespace subtype theorem tag_eq {a1 a2 : A} {H1 : P a1} {H2 : P a2} (H3 : a1 = a2) : tag a1 H1 = tag a2 H2 := eq.subst H3 (take H2, tag_irrelevant H1 H2) H2 - protected theorem equal {a1 a2 : {x | P x}} : ∀(H : elt_of a1 = elt_of a2), a1 = a2 := + protected theorem eq {a1 a2 : {x | P x}} : ∀(H : elt_of a1 = elt_of a2), a1 = a2 := destruct a1 (take x1 H1, destruct a2 (take x2 H2 H, tag_eq H)) protected definition is_inhabited [instance] {a : A} (H : P a) : inhabited {x | P x} := diff --git a/library/data/unit.lean b/library/data/unit.lean index 4d340112b..195bf79fe 100644 --- a/library/data/unit.lean +++ b/library/data/unit.lean @@ -8,18 +8,18 @@ import logic.eq namespace unit notation `⋆` := star - protected theorem equal (a b : unit) : a = b := + protected theorem eq (a b : unit) : a = b := unit.rec_on a (unit.rec_on b rfl) theorem eq_star (a : unit) : a = star := - unit.equal a star + unit.eq a star protected theorem subsingleton [instance] : subsingleton unit := - subsingleton.intro (λ a b, unit.equal a b) + subsingleton.intro (λ a b, unit.eq a b) protected definition is_inhabited [instance] : inhabited unit := inhabited.mk unit.star protected definition has_decidable_eq [instance] : decidable_eq unit := - take (a b : unit), decidable.inl (unit.equal a b) + take (a b : unit), decidable.inl (unit.eq a b) end unit diff --git a/library/init/quot.lean b/library/init/quot.lean index 5a283b47c..c95ec18b7 100644 --- a/library/init/quot.lean +++ b/library/init/quot.lean @@ -48,7 +48,7 @@ namespace quot protected lemma indep_coherent (f : Π a, B ⟦a⟧) (H : ∀ (a b : A) (p : a ≈ b), eq.rec (f a) (sound p) = f b) : ∀ a b, a ≈ b → quot.indep f a = quot.indep f b := - λa b e, sigma.equal (sound e) (H a b e) + λa b e, sigma.eq (sound e) (H a b e) protected lemma lift_indep_pr1 (f : Π a, B ⟦a⟧) (H : ∀ (a b : A) (p : a ≈ b), eq.rec (f a) (sound p) = f b) diff --git a/library/init/sigma.lean b/library/init/sigma.lean index 3a7f02601..b51764c3e 100644 --- a/library/init/sigma.lean +++ b/library/init/sigma.lean @@ -30,7 +30,7 @@ namespace sigma theorem dpair_eq : ∀ {a₁ a₂ : A} {b₁ : B a₁} {b₂ : B a₂} (H₁ : a₁ = a₂), eq.rec_on H₁ b₁ = b₂ → ⟨a₁, b₁⟩ = ⟨a₂, b₂⟩ | a₁ a₁ b₁ b₁ rfl rfl := rfl - protected theorem equal {p₁ p₂ : Σa : A, B a} : + protected theorem eq {p₁ p₂ : Σa : A, B a} : ∀(H₁ : p₁.1 = p₂.1) (H₂ : eq.rec_on H₁ p₁.2 = p₂.2), p₁ = p₂ := destruct p₁ (take a₁ b₁, destruct p₂ (take a₂ b₂ H₁ H₂, dpair_eq H₁ H₂))