refactoring of Isomorphsim complete

This commit is contained in:
wadler 2018-01-31 16:49:53 -02:00
parent e39a019602
commit 871573291a
2 changed files with 20 additions and 19 deletions

View file

@ -26,7 +26,7 @@ logic.
import Relation.Binary.PropositionalEquality as Eq import Relation.Binary.PropositionalEquality as Eq
open Eq using (_≡_; refl; sym; trans; cong) open Eq using (_≡_; refl; sym; trans; cong)
open Eq.≡-Reasoning open Eq.≡-Reasoning
open import Isomorphism using (_≃_; ≃-sym; _≲_) open import Isomorphism using (_≃_; ≃-sym; ≃-trans; _≲_)
open Isomorphism.≃-Reasoning open Isomorphism.≃-Reasoning
open import Data.Nat using (; zero; suc; _+_; _*_) open import Data.Nat using (; zero; suc; _+_; _*_)
open import Data.Nat.Properties.Simple using (+-suc) open import Data.Nat.Properties.Simple using (+-suc)
@ -257,22 +257,21 @@ same as `Bool`. But there is an isomorphism betwee the two types.
For instance, `(tt, true)`, which is a member of the former, For instance, `(tt, true)`, which is a member of the former,
corresponds to `true`, which is a member of the latter. corresponds to `true`, which is a member of the latter.
That unit is also a right identity follows immediately from left That unit is also a right identity follows immediately from
identity, commutativity of product, and symmetry and transitivity of commutativity of product and left identity.
isomorphism.
\begin{code} \begin{code}
-identityʳ : ∀ {A : Set} → A ≃ (A × ) -identityʳ : ∀ {A : Set} → (A × ) ≃ A
-identityʳ {A} = -identityʳ {A} =
≃-begin ≃-begin
(A × )
≃⟨ ×-comm ⟩
( × A)
≃⟨ -identityˡ ⟩
A A
≃⟨ ≃-sym -identityˡ ⟩
× A
≃⟨ ×-comm {} {A} ⟩
A ×
≃-∎ ≃-∎
\end{code} \end{code}
Here we have used tabular reasoning for isomorphism,
[TODO: We could introduce a notation similar to that for reasoning on ≡.] analogous to than used for equivalence.
## Disjunction is sum ## Disjunction is sum
@ -490,20 +489,21 @@ isomorphism betwee the two types. For instance, `inj₂ true`, which is
a member of the former, corresponds to `true`, which is a member of a member of the former, corresponds to `true`, which is a member of
the latter. the latter.
That empty is also a right identity follows immediately from left That empty is also a right identity follows immediately from
identity, commutativity of sum, and symmetry and transitivity of commutativity of sum and left identity.
isomorphism.
\begin{code} \begin{code}
⊥-identityʳ : ∀ {A : Set} → A ≃ (A ⊎ ⊥) ⊥-identityʳ : ∀ {A : Set} → (A ⊎ ⊥) ≃ A
⊥-identityʳ {A} = ⊥-identityʳ {A} =
≃-begin ≃-begin
A (A ⊎ ⊥)
≃⟨ sym ⊥-identityˡ ⟩
× A
≃⟨ ⊎-comm ⟩ ≃⟨ ⊎-comm ⟩
A × (⊥ ⊎ A)
≃⟨ ⊥-identityˡ ⟩
A
≃-∎ ≃-∎
\end{code} \end{code}
Here we have used tabular reasoning for isomorphism,
analogous to than used for equivalence.
## Implication is function ## Implication is function

View file

@ -6,6 +6,7 @@ permalink : /LogicAns
\begin{code} \begin{code}
open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym) open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym)
open import Isomorphism using (_≃_)
open import Logic open import Logic
\end{code} \end{code}