mirror of
https://github.com/Smaug123/agdaproofs
synced 2025-10-13 07:38:40 +00:00
Progress towards UFDs (#88)
This commit is contained in:
55
Maybe.agda
55
Maybe.agda
@@ -4,36 +4,41 @@ open import LogicalFormulae
|
||||
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
|
||||
|
||||
module Maybe where
|
||||
data Maybe {a : _} (A : Set a) : Set a where
|
||||
no : Maybe A
|
||||
yes : A → Maybe A
|
||||
|
||||
joinMaybe : {a : _} → {A : Set a} → Maybe (Maybe A) → Maybe A
|
||||
joinMaybe no = no
|
||||
joinMaybe (yes s) = s
|
||||
data Maybe {a : _} (A : Set a) : Set a where
|
||||
no : Maybe A
|
||||
yes : A → Maybe A
|
||||
|
||||
bindMaybe : {a b : _} → {A : Set a} → {B : Set b} → Maybe A → (A → Maybe B) → Maybe B
|
||||
bindMaybe no f = no
|
||||
bindMaybe (yes x) f = f x
|
||||
joinMaybe : {a : _} → {A : Set a} → Maybe (Maybe A) → Maybe A
|
||||
joinMaybe no = no
|
||||
joinMaybe (yes s) = s
|
||||
|
||||
applyMaybe : {a b : _} → {A : Set a} → {B : Set b} → Maybe (A → B) → Maybe A → Maybe B
|
||||
applyMaybe f no = no
|
||||
applyMaybe no (yes x) = no
|
||||
applyMaybe (yes f) (yes x) = yes (f x)
|
||||
bindMaybe : {a b : _} → {A : Set a} → {B : Set b} → Maybe A → (A → Maybe B) → Maybe B
|
||||
bindMaybe no f = no
|
||||
bindMaybe (yes x) f = f x
|
||||
|
||||
yesInjective : {a : _} → {A : Set a} → {x y : A} → (yes x ≡ yes y) → x ≡ y
|
||||
yesInjective {a} {A} {x} {.x} refl = refl
|
||||
applyMaybe : {a b : _} → {A : Set a} → {B : Set b} → Maybe (A → B) → Maybe A → Maybe B
|
||||
applyMaybe f no = no
|
||||
applyMaybe no (yes x) = no
|
||||
applyMaybe (yes f) (yes x) = yes (f x)
|
||||
|
||||
mapMaybe : {a b : _} → {A : Set a} → {B : Set b} → (f : A → B) → Maybe A → Maybe B
|
||||
mapMaybe f no = no
|
||||
mapMaybe f (yes x) = yes (f x)
|
||||
yesInjective : {a : _} → {A : Set a} → {x y : A} → (yes x ≡ yes y) → x ≡ y
|
||||
yesInjective {a} {A} {x} {.x} refl = refl
|
||||
|
||||
defaultValue : {a : _} → {A : Set a} → (default : A) → Maybe A → A
|
||||
defaultValue default no = default
|
||||
defaultValue default (yes x) = x
|
||||
mapMaybe : {a b : _} → {A : Set a} → {B : Set b} → (f : A → B) → Maybe A → Maybe B
|
||||
mapMaybe f no = no
|
||||
mapMaybe f (yes x) = yes (f x)
|
||||
|
||||
noNotYes : {a : _} {A : Set a} {b : A} → (no ≡ yes b) → False
|
||||
noNotYes ()
|
||||
defaultValue : {a : _} → {A : Set a} → (default : A) → Maybe A → A
|
||||
defaultValue default no = default
|
||||
defaultValue default (yes x) = x
|
||||
|
||||
mapMaybePreservesNo : {a b : _} {A : Set a} {B : Set b} {f : A → B} {x : Maybe A} → mapMaybe f x ≡ no → x ≡ no
|
||||
mapMaybePreservesNo {f = f} {no} pr = refl
|
||||
noNotYes : {a : _} {A : Set a} {b : A} → (no ≡ yes b) → False
|
||||
noNotYes ()
|
||||
|
||||
mapMaybePreservesNo : {a b : _} {A : Set a} {B : Set b} {f : A → B} {x : Maybe A} → mapMaybe f x ≡ no → x ≡ no
|
||||
mapMaybePreservesNo {f = f} {no} pr = refl
|
||||
|
||||
mapMaybePreservesYes : {a b : _} {A : Set a} {B : Set b} {f : A → B} {x : Maybe A} {y : B} → mapMaybe f x ≡ yes y → Sg A (λ z → (x ≡ yes z) && (f z ≡ y))
|
||||
mapMaybePreservesYes {f = f} {x} {y} map with x
|
||||
mapMaybePreservesYes {f = f} {x} {y} map | yes z = z , (refl ,, yesInjective map)
|
||||
|
Reference in New Issue
Block a user