annotate DPP/logic.agda @ 1:9f6cb9166d06

WIP dpp
author soto <soto@cr.ie.u-ryukyu.ac.jp>
date Sun, 01 May 2022 15:17:52 +0900
parents
children
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
1
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
1 module logic where
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
2
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
3 open import Level
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
4 open import Relation.Nullary
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
5 open import Relation.Binary hiding (_⇔_)
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
6 open import Relation.Binary.PropositionalEquality
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
7
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
8 open import Data.Empty
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
9 open import Data.Nat hiding (_⊔_)
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
10
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
11
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
12 data Bool : Set where
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
13 true : Bool
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
14 false : Bool
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
15
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
16 record _∧_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
17 constructor ⟪_,_⟫
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
18 field
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
19 proj1 : A
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
20 proj2 : B
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
21
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
22 data _∨_ {n m : Level} (A : Set n) ( B : Set m ) : Set (n ⊔ m) where
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
23 case1 : A → A ∨ B
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
24 case2 : B → A ∨ B
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
25
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
26 _⇔_ : {n m : Level } → ( A : Set n ) ( B : Set m ) → Set (n ⊔ m)
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
27 _⇔_ A B = ( A → B ) ∧ ( B → A )
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
28
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
29 contra-position : {n m : Level } {A : Set n} {B : Set m} → (A → B) → ¬ B → ¬ A
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
30 contra-position {n} {m} {A} {B} f ¬b a = ¬b ( f a )
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
31
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
32 double-neg : {n : Level } {A : Set n} → A → ¬ ¬ A
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
33 double-neg A notnot = notnot A
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
34
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
35 double-neg2 : {n : Level } {A : Set n} → ¬ ¬ ¬ A → ¬ A
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
36 double-neg2 notnot A = notnot ( double-neg A )
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
37
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
38 de-morgan : {n : Level } {A B : Set n} → A ∧ B → ¬ ( (¬ A ) ∨ (¬ B ) )
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
39 de-morgan {n} {A} {B} and (case1 ¬A) = ⊥-elim ( ¬A ( _∧_.proj1 and ))
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
40 de-morgan {n} {A} {B} and (case2 ¬B) = ⊥-elim ( ¬B ( _∧_.proj2 and ))
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
41
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
42 dont-or : {n m : Level} {A : Set n} { B : Set m } → A ∨ B → ¬ A → B
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
43 dont-or {A} {B} (case1 a) ¬A = ⊥-elim ( ¬A a )
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
44 dont-or {A} {B} (case2 b) ¬A = b
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
45
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
46 dont-orb : {n m : Level} {A : Set n} { B : Set m } → A ∨ B → ¬ B → A
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
47 dont-orb {A} {B} (case2 b) ¬B = ⊥-elim ( ¬B b )
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
48 dont-orb {A} {B} (case1 a) ¬B = a
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
49
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
50
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
51 infixr 130 _∧_
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
52 infixr 140 _∨_
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
53 infixr 150 _⇔_
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
54
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
55 _/\_ : Bool → Bool → Bool
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
56 true /\ true = true
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
57 _ /\ _ = false
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
58
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
59 _<B?_ : ℕ → ℕ → Bool
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
60 ℕ.zero <B? x = true
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
61 ℕ.suc x <B? ℕ.zero = false
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
62 ℕ.suc x <B? ℕ.suc xx = x <B? xx
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
63
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
64 -- _<BT_ : ℕ → ℕ → Set
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
65 -- ℕ.zero <BT ℕ.zero = ⊤
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
66 -- ℕ.zero <BT ℕ.suc b = ⊤
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
67 -- ℕ.suc a <BT ℕ.zero = ⊥
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
68 -- ℕ.suc a <BT ℕ.suc b = a <BT b
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
69
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
70
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
71 _≟B_ : Decidable {A = Bool} _≡_
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
72 true ≟B true = yes refl
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
73 false ≟B false = yes refl
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
74 true ≟B false = no λ()
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
75 false ≟B true = no λ()
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
76
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
77 _\/_ : Bool → Bool → Bool
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
78 false \/ false = false
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
79 _ \/ _ = true
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
80
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
81 not_ : Bool → Bool
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
82 not true = false
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
83 not false = true
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
84
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
85 _<=>_ : Bool → Bool → Bool
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
86 true <=> true = true
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
87 false <=> false = true
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
88 _ <=> _ = false
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
89
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
90 infixr 130 _\/_
9f6cb9166d06 WIP dpp
soto <soto@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
91 infixr 140 _/\_