module Realisability.PCA.Trivial where
Trivial PCAs🔗
The definition of partial combinatory algebra, either in terms of an abstraction elimination procedure or using a complete combinatorial basis, does not actually mandate that any elements are distinct. Therefore, the unit type can be equipped with the structure of a PCA.
: PCA lzero
⊤PCA .fst = el! ⊤
⊤PCA .snd .PCA-on.has-is-set = hlevel 2
⊤PCA .snd .PCA-on._%_ _ _ = pure tt ⊤PCA
To implement the abstraction procedure, we can simply map every term to the constant term containing the unit value.
.snd .PCA-on.has-is-pca = record where
⊤PCA open eval (λ _ _ → pure tt) renaming (eval to ⊤ev ; inst to ⊤inst) using ()
: ∀ {n} → Term ⊤ (suc n) → Term ⊤ n
abs _ = const (pure tt , tt)
abs
: ∀ {n} (t : Term ⊤ n) (ρ : Vec (↯⁺ ⊤) n) → ⌞ ⊤ev t ρ ⌟
rem = λ where
rem (var x) ρ → lookup ρ x .snd
(const x) ρ → x .snd
(app f x) ρ → tt
: {n : Nat} (t : Term ⊤ (suc n)) (ρ : Vec (↯⁺ ⊤) n) (a : ↯⁺ ⊤) → _
abs-β = part-ext
abs-β t ρ a (λ _ → rem (⊤inst t (const a)) ρ) (λ _ → tt) λ _ _ → refl
module _ {ℓ} (𝔸 : PCA ℓ) where
open Realisability.Data.Bool 𝔸
open Realisability.PCA.Sugar 𝔸
However, we can actually define what it means for a pca to be trivial
purely in terms of the programs it implements: if the `true
and `false
programs are identical, then the
pca is actually trivial. Following this, we define:
: Type _
is-trivial-pca = `true .fst ≡ `false .fst is-trivial-pca
A partial combinatory
algebra
is trivial when the programs `true
and `false
are identical in
this implies that the type
is a proposition.
: is-trivial-pca → is-prop ⌞ 𝔸 ⌟
is-trivial-pca→is-prop = always-injective $
is-trivial-pca→is-prop true=false x y
pure x ≡˘⟨ `true-β tt tt ⟩(λ e → e % pure x % pure y) true=false ⟩
`true ⋆ x ⋆ y ≡⟨ ap
`false ⋆ x ⋆ y ≡⟨ `false-β tt tt ⟩ pure y ∎
We define nontriviality to simply be the negation of triviality. Note that is nontrivial as soon as it contains two distinct programs, since if we are given then if were trivial in the sense above we would have which contradicts
: Type _
is-nontrivial-pca = `true .fst ≠ `false .fst
is-nontrivial-pca
: {x y : ⌞ 𝔸 ⌟} → x ≠ y → is-nontrivial-pca
two-elements→is-nontrivial = x≠y (is-trivial-pca→is-prop triv _ _) two-elements→is-nontrivial x≠y triv