2024-08-01 15:32:27 +00:00
|
|
|
|
open import Agda.Primitive
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
private
|
|
|
|
|
variable
|
|
|
|
|
l : Level
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
data _≡_ {A : Set l} : A → A → Set l where
|
|
|
|
|
refl : {x : A} → x ≡ x
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
data ⊥ : Set where
|
|
|
|
|
data ⊤ : Set where
|
|
|
|
|
tt : ⊤
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
result2 : (B : Set) → ((A : Set) → (B → A) → A) → B
|
|
|
|
|
result2 B x =
|
|
|
|
|
let
|
|
|
|
|
y = x B λ x → x
|
|
|
|
|
in y
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
data N : Set where
|
|
|
|
|
zero : N
|
|
|
|
|
suc : N -> N
|
|
|
|
|
{-# BUILTIN NATURAL N #-}
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
data Bool : Set where
|
|
|
|
|
true : Bool
|
|
|
|
|
false : Bool
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
ifbool : {A : Set} (x y : A) -> Bool -> A
|
|
|
|
|
ifbool {A} x y true = x
|
|
|
|
|
ifbool {A} x y false = y
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
negbool : Bool -> Bool
|
|
|
|
|
negbool true = false
|
|
|
|
|
negbool false = true
|
2024-07-29 22:45:35 +00:00
|
|
|
|
|
2024-08-01 15:32:27 +00:00
|
|
|
|
pred : N -> N
|
|
|
|
|
pred zero = zero
|
|
|
|
|
pred (suc x) = x
|
|
|
|
|
|
|
|
|
|
isZero : N -> Bool
|
|
|
|
|
isZero zero = true
|
|
|
|
|
isZero (suc x) = false
|
|
|
|
|
|
|
|
|
|
iter : (A : Set) (a : A) (f : A -> A) -> N -> A
|
|
|
|
|
iter A a f zero = a
|
|
|
|
|
iter A a f (suc x) = iter A (f a) f x
|
|
|
|
|
|
|
|
|
|
_^_ : {A : Set} → (A → A) → N → A → A
|
|
|
|
|
_^_ = λ f n → (λ x → iter _ x f n)
|
|
|
|
|
|
|
|
|
|
sub : N -> N -> N
|
|
|
|
|
sub x y = (pred ^ y) x
|
|
|
|
|
|
|
|
|
|
lt1 : N -> N -> Bool
|
|
|
|
|
lt1 m n = negbool (isZero (sub m n))
|
|
|
|
|
|
|
|
|
|
lt2 : N -> N -> Bool
|
|
|
|
|
lt2 m n = isZero (sub (suc m) n)
|
|
|
|
|
|
|
|
|
|
postulate
|
|
|
|
|
funExt : {A : Set} {B : A → Set} {f g : (x : A) → B x}
|
|
|
|
|
→ ((x : A) → f x ≡ g x) → f ≡ g
|
|
|
|
|
|
|
|
|
|
ap : {A B : Set l} (f : A → B) {x y : A} (p : x ≡ y) → f x ≡ f y
|
|
|
|
|
ap f refl = refl
|
|
|
|
|
|
|
|
|
|
trans : {A : Set l} {x y z : A} → x ≡ y → y ≡ z → x ≡ z
|
|
|
|
|
trans refl refl = refl
|
|
|
|
|
|
|
|
|
|
module ≡-Reasoning where
|
|
|
|
|
infix 1 begin_
|
|
|
|
|
begin_ : {l : Level} {A : Set l} {x y : A} → (x ≡ y) → (x ≡ y)
|
|
|
|
|
begin x = x
|
|
|
|
|
|
|
|
|
|
_≡⟨⟩_ : {l : Level} {A : Set l} (x {y} : A) → x ≡ y → x ≡ y
|
|
|
|
|
_ ≡⟨⟩ x≡y = x≡y
|
|
|
|
|
|
|
|
|
|
infixr 2 _≡⟨⟩_ step-≡
|
|
|
|
|
step-≡ : {l : Level} {A : Set l} (x {y z} : A) → y ≡ z → x ≡ y → x ≡ z
|
|
|
|
|
step-≡ _ y≡z x≡y = trans x≡y y≡z
|
|
|
|
|
syntax step-≡ x y≡z x≡y = x ≡⟨ x≡y ⟩ y≡z
|
|
|
|
|
|
|
|
|
|
infix 3 _∎
|
|
|
|
|
_∎ : {l : Level} {A : Set l} (x : A) → (x ≡ x)
|
|
|
|
|
_ ∎ = refl
|
|
|
|
|
open ≡-Reasoning
|
|
|
|
|
|
|
|
|
|
---
|
|
|
|
|
|
|
|
|
|
sub-zero-zero : (y : N) -> sub zero y ≡ zero
|
|
|
|
|
sub-zero-zero zero = refl
|
|
|
|
|
sub-zero-zero (suc y) = sub-zero-zero y
|
|
|
|
|
-- sub zero (suc y) ≡⟨⟩
|
|
|
|
|
-- (pred ^ (suc y)) zero ≡⟨⟩
|
|
|
|
|
-- iter _ zero pred (suc y) ≡⟨⟩
|
|
|
|
|
-- iter _ (pred zero) pred y ≡⟨⟩
|
|
|
|
|
-- iter _ zero pred y ≡⟨⟩
|
|
|
|
|
-- (pred ^ y) zero ≡⟨⟩
|
|
|
|
|
-- sub zero y ≡⟨ sub-zero-zero y ⟩
|
|
|
|
|
-- zero ∎
|
|
|
|
|
|
|
|
|
|
f : (x y : N) -> lt1 x y ≡ lt2 x y
|
|
|
|
|
f zero zero = refl
|
|
|
|
|
f zero (suc y) =
|
|
|
|
|
lt1 zero (suc y) ≡⟨⟩
|
|
|
|
|
negbool (isZero (sub zero (suc y))) ≡⟨ {! !} ⟩
|
|
|
|
|
lt2 zero (suc y) ∎
|
|
|
|
|
f (suc x) y = {! !}
|
|
|
|
|
|
|
|
|
|
prop : lt1 ≡ lt2
|
|
|
|
|
prop = funExt λ x → funExt λ y → f x y
|