Quantcast
Channel: September 2012 – Homotopy Type Theory
Viewing all articles
Browse latest Browse all 3

Isomorphism implies equality

$
0
0

Thierry Coquand and I have proved that, for a large class of algebraic structures, isomorphism implies equality (assuming univalence).

A class of algebraic structures

Structures in this class consist of a type, some operations on this type, and propositional axioms that can refer to operations and other axioms.

N-ary functions are defined in the following way (using Agda notation):

_^_⟶_ : Set → ℕ → Set → Set
A ^ zero ⟶ B = B
A ^ suc n ⟶ B = A → (A ^ n ⟶ B)

The class of algebraic structures is then defined in two steps. In the first step codes for structures, Structure, are defined mutually with a decoding function, ⟦_⟧, using induction-recursion:

mutual

  infixl 5 _+operator_ _+axiom_

  data Structure : Set₁ where
    empty : Structure

    -- N-ary functions.
    _+operator_ : Structure → (n : ℕ) → Structure

    -- Arbitrary propositional axioms.
    _+axiom_ : (s : Structure)
               (P : Σ (∀ A → ⟦ s ⟧ A → Set) λ P →
                      ∀ A s → Propositional (P A s)) →
               Structure

  ⟦_⟧ : Structure → Set → Set₁
  ⟦ empty                 ⟧ A = ↑ _ ⊤
  ⟦ s +operator n         ⟧ A = ⟦ s ⟧ A × (A ^ n ⟶ A)
  ⟦ s +axiom (P , P-prop) ⟧ A = Σ (⟦ s ⟧ A) (P A)

Here Propositional B means that B is of h-level 1, and ↑ _ ⊤ stands for the unit type, lifted to Set₁.

In the second step the decoding function’s Set argument is instantiated:

⟪_⟫ : Structure → Set₁
⟪ s ⟫ = Σ Set ⟦ s ⟧

As a simple example of an algebraic structure in this class we can define semigroups (over sets). The definition uses function extensionality to prove that the axioms are propositional. Note that the first axiom states that the underlying type is a set. This assumption is used to prove that the other axiom is propositional:

semigroup :
  ({A : Set} {B : A → Set} → Extensionality A B) →
  Structure
semigroup ext =
  empty

  +axiom
    ( (λ A _ → Is-set A)
    , is-set-prop
    )

  +operator 2

  +axiom
    ( (λ { _ (_ , _∙_) →
           ∀ x y z → x ∙ (y ∙ z) ≡ (x ∙ y) ∙ z })
    , assoc-prop
    )

  where
  is-set-prop = …

  assoc-prop = λ { _ ((_ , A-set) , _) → … }

Semigroup :
  ({A : Set} {B : A → Set} → Extensionality A B) →
  Set₁
Semigroup ext = ⟪ semigroup ext ⟫

Isomorphisms

We can also define what it means for two structures to be isomorphic. The property of being an n-ary function morphism is defined in the following way:

Is-_-ary-morphism :
  (n : ℕ) {A B : Set} → (A ^ n ⟶ A) → (B ^ n ⟶ B) → (A → B) → Set
Is- zero  -ary-morphism f₁ f₂ m = m f₁ ≡ f₂
Is- suc n -ary-morphism f₁ f₂ m =
  ∀ x → Is- n -ary-morphism (f₁ x) (f₂ (m x)) m

This definition can be lifted to structures by ignoring the (propositional) axioms:

Is-structure-morphism :
  (s : Structure) →
  {A B : Set} → ⟦ s ⟧ A → ⟦ s ⟧ B →
  (A → B) → Set
Is-structure-morphism empty           _          _          m = ⊤
Is-structure-morphism (s +axiom _)    (s₁ , _)   (s₂ , _)   m =
  Is-structure-morphism s s₁ s₂ m
Is-structure-morphism (s +operator n) (s₁ , op₁) (s₂ , op₂) m =
  Is-structure-morphism s s₁ s₂ m × Is- n -ary-morphism op₁ op₂ m

Two “top-level” structures are then defined to be isomorphic if there is a homomorphic bijection between the corresponding Sets:

Isomorphism : (s : Structure) → ⟪ s ⟫ → ⟪ s ⟫ → Set
Isomorphism s (A , s₁) (B , s₂) =
  Σ (A ↔ B) λ m → Is-structure-morphism s s₁ s₂ (_↔_.to m)

Here m : A ↔ B means that m is a bijection from A to B, and _↔_.to m is the corresponding function of type A → B.

The result

Finally we can state the main result. Isomorphic structures are equal, assuming univalence:

Univalence-axiom′ (Set ²/≡) Set →
Univalence-axiom lzero →
(s : Structure) (s₁ s₂ : ⟪ s ⟫) →
Isomorphism s s₁ s₂ → s₁ ≡ s₂

Here Univalence-axiom lzero is the full univalence axiom at universe level zero, and Univalence-axiom′ (Set ²/≡) Set is a special case of the univalence axiom at universe level one.

For the proof, see the code listing. The proof uses Voevodsky’s transport theorem (called subst-unique in the code).

In the future we may generalise the result to a larger class of structures, but the class above fits nicely in a blog post.


Viewing all articles
Browse latest Browse all 3

Trending Articles