refactor(library): rename namespace eq_ops to eq.ops

This commit is contained in:
Leonardo de Moura 2014-10-01 17:51:17 -07:00
parent ead827d6b7
commit 716cd4d651
37 changed files with 45 additions and 44 deletions

View file

@ -2,7 +2,7 @@
-- Released under Apache 2.0 license as described in the file LICENSE. -- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Leonardo de Moura -- Author: Leonardo de Moura
import logic.core.eq import logic.core.eq
open eq_ops open eq.ops
namespace binary namespace binary
context context

View file

@ -8,7 +8,7 @@ import data.unit data.sigma data.prod
import algebra.function import algebra.function
import logic.axioms.funext import logic.axioms.funext
open eq eq_ops open eq eq.ops
inductive category [class] (ob : Type) : Type := inductive category [class] (ob : Type) : Type :=
mk : Π (mor : ob → ob → Type) (comp : Π⦃A B C : ob⦄, mor B C → mor A B → mor A C) mk : Π (mor : ob → ob → Type) (comp : Π⦃A B C : ob⦄, mor B C → mor A B → mor A C)
@ -476,7 +476,7 @@ namespace category
end product end product
section arrow section arrow
open sigma eq_ops open sigma eq.ops
-- theorem concat_commutative_squares {ob : Type} {C : category ob} {a1 a2 a3 b1 b2 b3 : ob} -- theorem concat_commutative_squares {ob : Type} {C : category ob} {a1 a2 a3 b1 b2 b3 : ob}
-- {f1 : a1 => b1} {f2 : a2 => b2} {f3 : a3 => b3} {g2 : a2 => a3} {g1 : a1 => a2} -- {f1 : a1 => b1} {f2 : a2 => b2} {f3 : a3 => b3} {g2 : a2 => a3} {g1 : a1 => a2}
-- {h2 : b2 => b3} {h1 : b1 => b2} (H1 : f2 ∘ g1 = h1 ∘ f1) (H2 : f3 ∘ g2 = h2 ∘ f2) -- {h2 : b2 => b3} {h1 : b1 => b2} (H1 : f2 ∘ g1 = h1 ∘ f1) (H2 : f3 ∘ g2 = h2 ∘ f2)

View file

@ -4,7 +4,7 @@
import general_notation import general_notation
import logic.core.connectives logic.core.decidable logic.core.inhabited import logic.core.connectives logic.core.decidable logic.core.inhabited
open eq_ops eq decidable open eq eq.ops decidable
inductive bool : Type := inductive bool : Type :=
ff : bool, ff : bool,

View file

@ -14,7 +14,7 @@ import tools.fake_simplifier
open nat open nat
open quotient subtype prod relation open quotient subtype prod relation
open decidable binary fake_simplifier open decidable binary fake_simplifier
open eq_ops open eq.ops
namespace int namespace int
-- ## The defining equivalence relation on × -- ## The defining equivalence relation on ×

View file

@ -12,7 +12,7 @@ import .basic
open nat (hiding case) open nat (hiding case)
open decidable open decidable
open fake_simplifier open fake_simplifier
open int eq_ops open int eq.ops
namespace int namespace int

View file

@ -15,7 +15,7 @@ import logic tools.helper_tactics
import logic.core.identities import logic.core.identities
open nat open nat
open eq_ops open eq.ops
open helper_tactics open helper_tactics
inductive list (T : Type) : Type := inductive list (T : Type) : Type :=

View file

@ -10,7 +10,7 @@
import logic data.num tools.tactic algebra.binary tools.helper_tactics import logic data.num tools.tactic algebra.binary tools.helper_tactics
import logic.core.inhabited import logic.core.inhabited
open tactic binary eq_ops open tactic binary eq.ops
open decidable open decidable
open relation -- for subst_iff open relation -- for subst_iff
open helper_tactics open helper_tactics

View file

@ -13,7 +13,7 @@ import tools.fake_simplifier
open nat relation relation.iff_ops prod open nat relation relation.iff_ops prod
open fake_simplifier decidable open fake_simplifier decidable
open eq_ops open eq.ops
namespace nat namespace nat

View file

@ -10,7 +10,7 @@
import .basic logic.core.decidable import .basic logic.core.decidable
import tools.fake_simplifier import tools.fake_simplifier
open nat eq_ops tactic open nat eq.ops tactic
open fake_simplifier open fake_simplifier
namespace nat namespace nat

View file

@ -10,7 +10,7 @@
import data.nat.order import data.nat.order
import tools.fake_simplifier import tools.fake_simplifier
open nat eq_ops tactic open nat eq.ops tactic
open helper_tactics open helper_tactics
open fake_simplifier open fake_simplifier

View file

@ -3,7 +3,7 @@
-- Author: Leonardo de Moura -- Author: Leonardo de Moura
import logic.core.eq logic.core.inhabited logic.core.decidable import logic.core.eq logic.core.inhabited logic.core.decidable
open eq_ops decidable open eq.ops decidable
inductive option (A : Type) : Type := inductive option (A : Type) : Type :=
none {} : option A, none {} : option A,

View file

@ -37,7 +37,7 @@ section
theorem prod_ext (p : prod A B) : pair (pr1 p) (pr2 p) = p := theorem prod_ext (p : prod A B) : pair (pr1 p) (pr2 p) = p :=
destruct p (λx y, eq.refl (x, y)) destruct p (λx y, eq.refl (x, y))
open eq_ops open eq.ops
theorem pair_eq {a1 a2 : A} {b1 b2 : B} (H1 : a1 = a2) (H2 : b1 = b2) : (a1, b1) = (a2, b2) := theorem pair_eq {a1 a2 : A} {b1 b2 : B} (H1 : a1 = a2) (H2 : b1 = b2) : (a1, b1) = (a2, b2) :=
H1 ▸ H2 ▸ rfl H1 ▸ H2 ▸ rfl

View file

@ -9,7 +9,7 @@ import logic tools.tactic ..subtype logic.core.cast algebra.relation data.prod
import logic.core.instances import logic.core.instances
import .util import .util
open relation prod inhabited nonempty tactic eq_ops open relation prod inhabited nonempty tactic eq.ops
open subtype relation.iff_ops open subtype relation.iff_ops
namespace quotient namespace quotient
@ -225,7 +225,7 @@ theorem image_tag {A B : Type} {f : A → B} (u : image f) : ∃a H, tag (f a) H
obtain a (H : fun_image f a = u), from fun_image_surj u, obtain a (H : fun_image f a = u), from fun_image_surj u,
exists_intro a (exists_intro (exists_intro a rfl) H) exists_intro a (exists_intro (exists_intro a rfl) H)
open eq_ops open eq.ops
theorem fun_image_eq {A B : Type} (f : A → B) (a a' : A) theorem fun_image_eq {A B : Type} (f : A → B) (a a' : A)
: (f a = f a') ↔ (fun_image f a = fun_image f a') := : (f a = f a') ↔ (fun_image f a = fun_image f a') :=

View file

@ -5,7 +5,7 @@
import logic ..prod algebra.relation import logic ..prod algebra.relation
import tools.fake_simplifier import tools.fake_simplifier
open prod eq_ops open prod eq.ops
open fake_simplifier open fake_simplifier
namespace quotient namespace quotient

View file

@ -4,7 +4,7 @@
--- Author: Jeremy Avigad, Leonardo de Moura --- Author: Jeremy Avigad, Leonardo de Moura
---------------------------------------------------------------------------------------------------- ----------------------------------------------------------------------------------------------------
import data.bool import data.bool
open eq_ops bool open eq.ops bool
namespace set namespace set
definition set (T : Type) := definition set (T : Type) :=

View file

@ -2,7 +2,7 @@
-- Released under Apache 2.0 license as described in the file LICENSE. -- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Leonardo de Moura, Jeremy Avigad, Floris van Doorn -- Author: Leonardo de Moura, Jeremy Avigad, Floris van Doorn
import logic.core.inhabited logic.core.eq import logic.core.inhabited logic.core.eq
open inhabited eq_ops open inhabited eq.ops
inductive sigma {A : Type} (B : A → Type) : Type := inductive sigma {A : Type} (B : A → Type) : Type :=
dpair : Πx : A, B x → sigma B dpair : Πx : A, B x → sigma B

View file

@ -2,7 +2,7 @@
-- Released under Apache 2.0 license as described in the file LICENSE. -- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Leonardo de Moura, Jeremy Avigad -- Author: Leonardo de Moura, Jeremy Avigad
import logic.core.prop logic.core.inhabited logic.core.decidable import logic.core.prop logic.core.inhabited logic.core.decidable
open inhabited decidable eq_ops open inhabited decidable eq.ops
-- data.sum -- data.sum
-- ======== -- ========
-- The sum type, aka disjoint union. -- The sum type, aka disjoint union.

View file

@ -2,7 +2,7 @@
-- Released under Apache 2.0 license as described in the file LICENSE. -- Released under Apache 2.0 license as described in the file LICENSE.
-- Author: Floris van Doorn -- Author: Floris van Doorn
import data.nat.basic data.empty import data.nat.basic data.empty
open nat eq_ops open nat eq.ops
inductive vector (T : Type) : → Type := inductive vector (T : Type) : → Type :=
nil {} : vector T 0, nil {} : vector T 0,

View file

@ -7,7 +7,7 @@
import logic.core.quantifiers logic.core.cast algebra.relation import logic.core.quantifiers logic.core.cast algebra.relation
open eq_ops open eq.ops
axiom prop_complete (a : Prop) : a = true a = false axiom prop_complete (a : Prop) : a = true a = false

View file

@ -3,7 +3,7 @@
-- Author: Leonardo de Moura -- Author: Leonardo de Moura
import logic.axioms.hilbert logic.axioms.funext import logic.axioms.hilbert logic.axioms.funext
open eq_ops nonempty inhabited open eq.ops nonempty inhabited
-- Diaconescus theorem -- Diaconescus theorem
-- Show that Excluded middle follows from -- Show that Excluded middle follows from

View file

@ -7,7 +7,7 @@
import .eq .quantifiers import .eq .quantifiers
open eq_ops open eq.ops
definition cast {A B : Type} (H : A = B) (a : A) : B := definition cast {A B : Type} (H : A = B) (a : A) : B :=
eq.rec a H eq.rec a H

View file

@ -155,7 +155,7 @@ end iff
calc_refl iff.refl calc_refl iff.refl
calc_trans iff.trans calc_trans iff.trans
open eq_ops open eq.ops
theorem eq_to_iff {a b : Prop} (H : a = b) : a ↔ b := theorem eq_to_iff {a b : Prop} (H : a = b) : a ↔ b :=
iff.intro (λ Ha, H ▸ Ha) (λ Hb, H⁻¹ ▸ Hb) iff.intro (λ Ha, H ▸ Ha) (λ Hb, H⁻¹ ▸ Hb)

View file

@ -36,18 +36,19 @@ namespace eq
theorem symm {A : Type} {a b : A} (H : a = b) : b = a := theorem symm {A : Type} {a b : A} (H : a = b) : b = a :=
subst H (refl a) subst H (refl a)
namespace ops
postfix `⁻¹` := symm
infixr `⬝` := trans
infixr `▸` := subst
end ops
end eq end eq
calc_subst eq.subst calc_subst eq.subst
calc_refl eq.refl calc_refl eq.refl
calc_trans eq.trans calc_trans eq.trans
namespace eq_ops open eq.ops
postfix `⁻¹` := eq.symm
infixr `⬝` := eq.trans
infixr `▸` := eq.subst
end eq_ops
open eq_ops
namespace eq namespace eq
-- eq_rec with arguments swapped, for transporting an element of a dependent type -- eq_rec with arguments swapped, for transporting an element of a dependent type

View file

@ -7,7 +7,7 @@ import ..instances
open relation open relation
open relation.general_operations open relation.general_operations
open relation.iff_ops open relation.iff_ops
open eq_ops open eq.ops
section section

View file

@ -5,7 +5,7 @@
---------------------------------------------------------------------------------------------------- ----------------------------------------------------------------------------------------------------
import logic.core.decidable tools.tactic import logic.core.decidable tools.tactic
open decidable tactic eq_ops open decidable tactic eq.ops
definition ite (c : Prop) {H : decidable c} {A : Type} (t e : A) : A := definition ite (c : Prop) {H : decidable c} {A : Type} (t e : A) : A :=
decidable.rec_on H (assume Hc, t) (assume Hnc, e) decidable.rec_on H (assume Hc, t) (assume Hnc, e)

View file

@ -1,5 +1,5 @@
import logic import logic
open bool eq_ops tactic eq open bool eq.ops tactic eq
variables a b c : bool variables a b c : bool
axiom H1 : a = b axiom H1 : a = b

View file

@ -3,7 +3,7 @@ import tools.fake_simplifier
open nat relation relation.iff_ops prod open nat relation relation.iff_ops prod
open fake_simplifier decidable open fake_simplifier decidable
open eq_ops open eq.ops
namespace nat namespace nat

View file

@ -10,7 +10,7 @@
-- Basic properties of lists. -- Basic properties of lists.
import data.nat import data.nat
open nat eq_ops open nat eq.ops
inductive list (T : Type) : Type := inductive list (T : Type) : Type :=
nil {} : list T, nil {} : list T,
cons : T → list T → list T cons : T → list T → list T

View file

@ -1,5 +1,5 @@
import logic import logic
open eq_ops open eq.ops
inductive nat : Type := inductive nat : Type :=
zero : nat, zero : nat,

View file

@ -1,5 +1,5 @@
import logic import logic
open eq_ops open eq.ops
inductive nat : Type := inductive nat : Type :=
zero : nat, zero : nat,

View file

@ -1,5 +1,5 @@
import logic import logic
open eq_ops open eq.ops
inductive nat : Type := inductive nat : Type :=
zero : nat, zero : nat,

View file

@ -1,5 +1,5 @@
import logic import logic
open eq_ops eq open eq.ops eq
inductive nat : Type := inductive nat : Type :=
zero : nat, zero : nat,

View file

@ -1,5 +1,5 @@
import logic import logic
open eq_ops open eq.ops
inductive nat : Type := inductive nat : Type :=
zero : nat, zero : nat,

View file

@ -1,5 +1,5 @@
import logic import logic
open bool eq_ops tactic open bool eq.ops tactic
variables a b c : bool variables a b c : bool
axiom H1 : a = b axiom H1 : a = b

View file

@ -14,7 +14,7 @@ import logic data.nat
open nat open nat
-- open congr -- open congr
open eq_ops eq open eq.ops eq
inductive list (T : Type) : Type := inductive list (T : Type) : Type :=
nil {} : list T, nil {} : list T,

View file

@ -4,7 +4,7 @@
-- Author: Floris van Doorn -- Author: Floris van Doorn
---------------------------------------------------------------------------------------------------- ----------------------------------------------------------------------------------------------------
import logic algebra.binary import logic algebra.binary
open tactic binary eq_ops eq open tactic binary eq.ops eq
open decidable open decidable
definition refl := @eq.refl definition refl := @eq.refl

View file

@ -4,7 +4,7 @@
-- Author: Floris van Doorn -- Author: Floris van Doorn
---------------------------------------------------------------------------------------------------- ----------------------------------------------------------------------------------------------------
import logic algebra.binary import logic algebra.binary
open tactic binary eq_ops eq open tactic binary eq.ops eq
open decidable open decidable
inductive nat : Type := inductive nat : Type :=