Banish qualified import of Function - use \o for fun-comp!

This commit is contained in:
Frederik Hanghøj Iversen 2018-04-11 11:10:33 +02:00
parent c90b064bb0
commit db5fb3603a
10 changed files with 29 additions and 43 deletions

View file

@ -7,7 +7,6 @@ open import Data.Bool hiding (T)
open import Data.Sum hiding ([_,_]) open import Data.Sum hiding ([_,_])
open import Data.Unit open import Data.Unit
open import Data.Empty open import Data.Empty
open import Function
open import Relation.Nullary open import Relation.Nullary
open import Relation.Nullary.Decidable open import Relation.Nullary.Decidable

View file

@ -2,7 +2,6 @@
module Cat.Categories.Fam where module Cat.Categories.Fam where
open import Cat.Prelude open import Cat.Prelude
import Function
open import Cat.Category open import Cat.Category
@ -15,7 +14,7 @@ module _ (a b : Level) where
fst identity = λ x x fst identity = λ x x
snd identity = λ b b snd identity = λ b b
_<<<_ : {a b c : Object} Arr b c Arr a b Arr a c _<<<_ : {a b c : Object} Arr b c Arr a b Arr a c
(g , g') <<< (f , f') = g Function. f , g' Function. f' (g , g') <<< (f , f') = g f , g' f'
RawFam : RawCategory (lsuc (a b)) (a b) RawFam : RawCategory (lsuc (a b)) (a b)
RawFam = record RawFam = record

View file

@ -2,7 +2,6 @@
module Cat.Categories.Rel where module Cat.Categories.Rel where
open import Cat.Prelude hiding (Rel) open import Cat.Prelude hiding (Rel)
open import Function
open import Cat.Category open import Cat.Category

View file

@ -4,8 +4,6 @@ module Cat.Categories.Sets where
open import Cat.Prelude as P open import Cat.Prelude as P
open import Function using (_∘_ ; _∘_)
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
open import Cat.Category.Product open import Cat.Category.Product
@ -25,14 +23,14 @@ module _ ( : Level) where
SetsRaw : RawCategory (lsuc ) SetsRaw : RawCategory (lsuc )
RawCategory.Object SetsRaw = hSet RawCategory.Object SetsRaw = hSet
RawCategory.Arrow SetsRaw (T , _) (U , _) = T U RawCategory.Arrow SetsRaw (T , _) (U , _) = T U
RawCategory.identity SetsRaw = Function.id RawCategory.identity SetsRaw = idFun _
RawCategory._<<<_ SetsRaw = Function._∘_ RawCategory._<<<_ SetsRaw = _∘_
module _ where module _ where
private private
open RawCategory SetsRaw hiding (_<<<_) open RawCategory SetsRaw hiding (_<<<_)
isIdentity : IsIdentity Function.id isIdentity : IsIdentity (idFun _)
fst isIdentity = funExt λ _ refl fst isIdentity = funExt λ _ refl
snd isIdentity = funExt λ _ refl snd isIdentity = funExt λ _ refl
@ -70,17 +68,17 @@ module _ ( : Level) where
where where
module xxA = AreInverses xx module xxA = AreInverses xx
module yyA = AreInverses yy module yyA = AreInverses yy
ve-re : g f Function.id ve-re : g f idFun _
ve-re = arrowsAreSets {A = hA} {B = hA} _ _ xxA.verso-recto yyA.verso-recto i ve-re = arrowsAreSets {A = hA} {B = hA} _ _ xxA.verso-recto yyA.verso-recto i
re-ve : f g Function.id re-ve : f g idFun _
re-ve = arrowsAreSets {A = hB} {B = hB} _ _ xxA.recto-verso yyA.recto-verso i re-ve = arrowsAreSets {A = hB} {B = hB} _ _ xxA.recto-verso yyA.recto-verso i
1eq : x.inverse y.inverse 1eq : x.inverse y.inverse
1eq = begin 1eq = begin
x.inverse ≡⟨⟩ x.inverse ≡⟨⟩
x.inverse Function.id ≡⟨ cong (λ φ x.inverse φ) (sym yA.recto-verso) x.inverse idFun _ ≡⟨ cong (λ φ x.inverse φ) (sym yA.recto-verso)
x.inverse (f y.inverse) ≡⟨⟩ x.inverse (f y.inverse) ≡⟨⟩
(x.inverse f) y.inverse ≡⟨ cong (λ φ φ y.inverse) xA.verso-recto (x.inverse f) y.inverse ≡⟨ cong (λ φ φ y.inverse) xA.verso-recto
Function.id y.inverse ≡⟨⟩ idFun _ y.inverse ≡⟨⟩
y.inverse y.inverse
2eq : (λ i AreInverses f (1eq i)) [ x.areInverses y.areInverses ] 2eq : (λ i AreInverses f (1eq i)) [ x.areInverses y.areInverses ]
2eq = lemPropF p 1eq 2eq = lemPropF p 1eq
@ -266,7 +264,7 @@ module _ { : Level} where
module _ (hX : Object) where module _ (hX : Object) where
open Σ hX renaming (fst to X) open Σ hX renaming (fst to X)
module _ (f : X A ) (g : X B) where module _ (f : X A ) (g : X B) where
ump : fst Function.∘′ (f &&& g) f × snd Function.∘′ (f &&& g) g ump : fst ∘′ (f &&& g) f × snd ∘′ (f &&& g) g
fst ump = refl fst ump = refl
snd ump = refl snd ump = refl
@ -280,7 +278,7 @@ module _ { : Level} where
= f &&& g , ump hX f g , λ eq funExt (umpUniq eq) = f &&& g , ump hX f g , λ eq funExt (umpUniq eq)
where where
open Σ hX renaming (fst to X) using () open Σ hX renaming (fst to X) using ()
module _ {y : X A × B} (eq : fst Function.∘′ y f × snd Function.∘′ y g) (x : X) where module _ {y : X A × B} (eq : fst ∘′ y f × snd ∘′ y g) (x : X) where
p1 : fst ((f &&& g) x) fst (y x) p1 : fst ((f &&& g) x) fst (y x)
p1 = begin p1 = begin
fst ((f &&& g) x) ≡⟨⟩ fst ((f &&& g) x) ≡⟨⟩

View file

@ -35,8 +35,6 @@ open Cat.Equivalence
renaming (_≅_ to _≈_) renaming (_≅_ to _≈_)
hiding (preorder≅ ; Isomorphism) hiding (preorder≅ ; Isomorphism)
import Function
------------------ ------------------
-- * Categories -- -- * Categories --
------------------ ------------------
@ -146,7 +144,7 @@ record RawCategory (a b : Level) : Set (lsuc (a ⊔ b)) where
postulate from[Contr] : Univalent[Contr] Univalent postulate from[Contr] : Univalent[Contr] Univalent
from[Andrea] : Univalent[Andrea] Univalent from[Andrea] : Univalent[Andrea] Univalent
from[Andrea] = from[Contr] Function. step from[Andrea] = from[Contr] step
where where
module _ (f : Univalent[Andrea]) (A : Object) where module _ (f : Univalent[Andrea]) (A : Object) where
aux : isContr (Σ[ B Object ] A B) aux : isContr (Σ[ B Object ] A B)
@ -510,7 +508,7 @@ module Opposite {a b : Level} where
module = Category module = Category
opRaw : RawCategory a b opRaw : RawCategory a b
RawCategory.Object opRaw = .Object RawCategory.Object opRaw = .Object
RawCategory.Arrow opRaw = Function.flip .Arrow RawCategory.Arrow opRaw = flip .Arrow
RawCategory.identity opRaw = .identity RawCategory.identity opRaw = .identity
RawCategory._<<<_ opRaw = ._>>>_ RawCategory._<<<_ opRaw = ._>>>_
@ -529,9 +527,6 @@ module Opposite {a b : Level} where
open Σ k renaming (fst to η ; snd to inv-η) open Σ k renaming (fst to η ; snd to inv-η)
open AreInverses inv-η open AreInverses inv-η
_⊙_ = Function._∘_
infixr 9 _⊙_
genericly : {a b c : Level} {a : Set a} {b : Set b} {c : Set c} genericly : {a b c : Level} {a : Set a} {b : Set b} {c : Set c}
a × b × c b × a × c a × b × c b × a × c
genericly (a , b , c) = (b , a , c) genericly (a , b , c) = (b , a , c)
@ -557,22 +552,22 @@ module Opposite {a b : Level} where
open Σ r-iso renaming (fst to r-l ; snd to r-r) open Σ r-iso renaming (fst to r-l ; snd to r-r)
ζ : A B A B ζ : A B A B
ζ = η shuffle ζ = η shuffle
-- inv : AreInverses (.idToIso A B) f -- inv : AreInverses (.idToIso A B) f
inv-ζ : AreInverses (idToIso A B) ζ inv-ζ : AreInverses (idToIso A B) ζ
-- recto-verso : .idToIso A B <<< f ≡ idFun (A .≅ B) -- recto-verso : .idToIso A B <<< f ≡ idFun (A .≅ B)
inv-ζ = record inv-ζ = record
{ verso-recto = funExt (λ x begin { verso-recto = funExt (λ x begin
(ζ idToIso A B) x ≡⟨⟩ (ζ idToIso A B) x ≡⟨⟩
(η shuffle idToIso A B) x ≡⟨ cong (λ φ φ x) (cong (λ φ η shuffle φ) (funExt lem)) (η shuffle idToIso A B) x ≡⟨ cong (λ φ φ x) (cong (λ φ η shuffle φ) (funExt lem))
(η shuffle shuffle~ .idToIso A B) x ≡⟨⟩ (η shuffle shuffle~ .idToIso A B) x ≡⟨⟩
(η .idToIso A B) x ≡⟨ (λ i verso-recto i x) (η .idToIso A B) x ≡⟨ (λ i verso-recto i x)
x ) x )
; recto-verso = funExt (λ x begin ; recto-verso = funExt (λ x begin
(idToIso A B η shuffle) x ≡⟨ cong (λ φ φ x) (cong (λ φ φ η shuffle) (funExt lem)) (idToIso A B η shuffle) x ≡⟨ cong (λ φ φ x) (cong (λ φ φ η shuffle) (funExt lem))
(shuffle~ .idToIso A B η shuffle) x ≡⟨ cong (λ φ φ x) (cong (λ φ shuffle~ φ shuffle) recto-verso) (shuffle~ .idToIso A B η shuffle) x ≡⟨ cong (λ φ φ x) (cong (λ φ shuffle~ φ shuffle) recto-verso)
(shuffle~ shuffle) x ≡⟨⟩ (shuffle~ shuffle) x ≡⟨⟩
x ) x )
} }

View file

@ -2,7 +2,6 @@
module Cat.Category.Functor where module Cat.Category.Functor where
open import Cat.Prelude open import Cat.Prelude
open import Function
open import Cubical open import Cubical
@ -165,8 +164,8 @@ module Functors where
module _ {c cc : Level} { : Category c cc} where module _ {c cc : Level} { : Category c cc} where
private private
raw : RawFunctor raw : RawFunctor
RawFunctor.omap raw = Function.id RawFunctor.omap raw = idFun _
RawFunctor.fmap raw = Function.id RawFunctor.fmap raw = idFun _
isFunctor : IsFunctor raw isFunctor : IsFunctor raw
IsFunctor.isIdentity isFunctor = refl IsFunctor.isIdentity isFunctor = refl

View file

@ -5,7 +5,6 @@ This module provides construction 2.3 in [voe]
module Cat.Category.Monad.Voevodsky where module Cat.Category.Monad.Voevodsky where
open import Cat.Prelude open import Cat.Prelude
open import Function
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor as F open import Cat.Category.Functor as F
@ -132,10 +131,10 @@ module voe {a b : Level} ( : Category a b) where
Kleisli→Monoidal : K.Monad M.Monad Kleisli→Monoidal : K.Monad M.Monad
Kleisli→Monoidal = E.reverse Kleisli→Monoidal = E.reverse
ve-re : Kleisli→Monoidal Monoidal→Kleisli Function.id ve-re : Kleisli→Monoidal Monoidal→Kleisli idFun _
ve-re = E.verso-recto ve-re = E.verso-recto
re-ve : Monoidal→Kleisli Kleisli→Monoidal Function.id re-ve : Monoidal→Kleisli Kleisli→Monoidal idFun _
re-ve = E.recto-verso re-ve = E.recto-verso
forth : §2-3.§1 omap pure §2-3.§2 omap pure forth : §2-3.§1 omap pure §2-3.§2 omap pure

View file

@ -3,7 +3,6 @@ module Cat.Category.Product where
open import Cat.Prelude as P hiding (_×_ ; fst ; snd) open import Cat.Prelude as P hiding (_×_ ; fst ; snd)
open import Cat.Equivalence hiding (_≅_) open import Cat.Equivalence hiding (_≅_)
-- module P = Cat.Prelude
open import Cat.Category open import Cat.Category
@ -203,15 +202,13 @@ module Try0 {a b : Level} { : Category a b}
: ((X , xa , xb) (Y , ya , yb)) : ((X , xa , xb) (Y , ya , yb))
(Σ[ p (X Y) ] (PathP (λ i .Arrow (p i) A) xa ya) × (PathP (λ i .Arrow (p i) B) xb yb)) (Σ[ p (X Y) ] (PathP (λ i .Arrow (p i) A) xa ya) × (PathP (λ i .Arrow (p i) B) xb yb))
step0 step0
= (λ p cong fst p , cong-d (fst snd) p , cong-d (snd snd) p) = (λ p cong fst p , cong-d (fst snd) p , cong-d (snd snd) p)
-- , (λ x → λ i → fst x i , (fst (snd x) i) , (snd (snd x) i)) -- , (λ x → λ i → fst x i , (fst (snd x) i) , (snd (snd x) i))
, (λ{ (p , q , r) Σ≡ p λ i q i , r i}) , (λ{ (p , q , r) Σ≡ p λ i q i , r i})
, record , record
{ verso-recto = funExt (λ{ p refl}) { verso-recto = funExt (λ{ p refl})
; recto-verso = funExt (λ{ (p , q , r) refl}) ; recto-verso = funExt (λ{ (p , q , r) refl})
} }
where
open import Function renaming (_∘_ to _⊙_)
-- Should follow from c being univalent -- Should follow from c being univalent
iso-id-inv : {p : X Y} p .isoToId (.idToIso X Y p) iso-id-inv : {p : X Y} p .isoToId (.idToIso X Y p)

View file

@ -151,8 +151,6 @@ module _ {a b : Level} (A : Set a) (B : Set b) where
(x f) y ≡⟨ cong (λ φ φ y) inv-x.verso-recto (x f) y ≡⟨ cong (λ φ φ y) inv-x.verso-recto
y y
open import Cat.Prelude
propInv : g isProp (AreInverses f g) propInv : g isProp (AreInverses f g)
propInv g t u i = record { verso-recto = a i ; recto-verso = b i } propInv g t u i = record { verso-recto = a i ; recto-verso = b i }
where where

View file

@ -9,7 +9,10 @@ open import Data.Product public
renaming (∃! to ∃!≈) renaming (∃! to ∃!≈)
using (_×_ ; Σ-syntax ; swap) using (_×_ ; Σ-syntax ; swap)
-- TODO Import Data.Function under appropriate names. open import Function using (_∘_ ; _∘_ ; _$_ ; case_of_ ; flip) public
idFun : {a} (A : Set a) A A
idFun A a = a
open import Cubical public open import Cubical public
-- FIXME rename `gradLemma` to `fromIsomorphism` - perhaps just use wrapper -- FIXME rename `gradLemma` to `fromIsomorphism` - perhaps just use wrapper