Files
agda-doodles/Peano.agda
2019-03-28 21:45:04 +01:00

87 lines
2.6 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module peano where
data : Set where
zero :
suc :
_+_ :
zero + zero = zero
zero + n = n
(suc n) + n = suc (n + n)
_*_ :
zero * n = zero
(suc zero) * n = n
(suc n) * n = n + (n * n)
data _even : Set where
ZERO : zero even
STEP : x x even suc (suc x) even
-- alternative:
-- : (x : ) x even suc (suc x) even
proof₁ : suc (suc (suc (suc zero))) even
proof₁ = STEP _ (STEP _ ZERO)
data _∧_ (P : Set) (Q : Set) : Set where
∧-intro : P Q (P Q)
proof₃ : {P Q : Set} (P Q) P
proof₃ (∧-intro p q) = p
_⇔_ : (P : Set) (Q : Set) Set
a b = (a b) (b a)
∧-comm : {P Q : Set} (P Q) (Q P)
∧-comm (∧-intro p q) = ∧-intro q p
--∧-comm : {P Q : Set} (P Q) (Q P)
--∧-comm = ∧-intro (∧-comm {P} {Q}) (∧-comm {Q} {P})
∧-assoc₁ : {P Q R : Set} ((P Q) R) (P (Q R))
∧-assoc₁ (∧-intro (∧-intro p q) r) = ∧-intro p (∧-intro q r)
∧-assoc₂ : {P Q R : Set} (P (Q R)) ((P Q) R)
∧-assoc₂ (∧-intro p (∧-intro q r)) = ∧-intro (∧-intro p q) r
∧-assoc : {P Q R : Set} ((P Q) R) (P (Q R))
∧-assoc = ∧-intro ∧-assoc₁ ∧-assoc₂
data __ (P Q : Set) : Set where
-intro₁ : P P Q
-intro₂ : Q P Q
-elim : {A B C : Set} (A C) (B C) (A B) C
-elim ac bc (-intro₁ a) = ac a
-elim ac bc (-intro₂ b) = bc b
-comm : {P Q : Set} (P Q) (Q P)
-comm (-intro₁ p) = -intro₂ p
-comm (-intro₂ q) = -intro₁ q
-comm : {P Q : Set} (P Q) (Q P)
-comm = ∧-intro -comm -comm
-- this is my first own proof! The associativity of disjunction
-assoc₁ : {P Q R : Set} ((P Q) R) (P (Q R))
-assoc₁ (-intro₁ (-intro₁ p)) = -intro₁ p
-assoc₁ (-intro₁ (-intro₂ q)) = -intro₂ (-intro₁ q)
-assoc₁ (-intro₂ r) = -intro₂ (-intro₂ r)
-assoc₂ : {P Q R : Set} (P (Q R)) ((P Q) R)
-assoc₂ (-intro₁ p) = -intro₁ (-intro₁ p)
-assoc₂ (-intro₂ (-intro₁ q)) = -intro₁ (-intro₂ q)
-assoc₂ (-intro₂ (-intro₂ r)) = -intro₂ r
-assoc : {P Q R : Set} ((P Q) R) (P (Q R))
-assoc = ∧-intro -assoc₁ -assoc₂
--
data : Set where -- nothing
¬ : Set Set
¬ A = A