annotate utilities.agda @ 66:9071e5a77a13

implies
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Mon, 23 Dec 2019 09:48:54 +0900
parents a39a82820742
children 52d957db0222
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
27
a39a82820742 add whileTestCondition
ryokka
parents: 21
diff changeset
1 {-# OPTIONS --allow-unsolved-metas #-}
12
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
2 module utilities where
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
3
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
4 open import Function
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
5 open import Data.Nat
21
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
6 open import Data.Product
27
a39a82820742 add whileTestCondition
ryokka
parents: 21
diff changeset
7 open import Data.Bool hiding ( _≟_ ; _≤?_)
12
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
8 open import Level renaming ( suc to succ ; zero to Zero )
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
9 open import Relation.Nullary using (¬_; Dec; yes; no)
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
10 open import Relation.Binary.PropositionalEquality
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
11
21
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
12 Pred : Set -> Set₁
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
13 Pred X = X -> Set
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
14
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
15 Imply : Set -> Set -> Set
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
16 Imply X Y = X -> Y
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
17
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
18 Iff : Set -> Set -> Set
5e4abc1919b4 fix module relation
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 15
diff changeset
19 Iff X Y = Imply X Y × Imply Y X
12
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
20
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
21 record _/\_ {n : Level } (a : Set n) (b : Set n): Set n where
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
22 field
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
23 pi1 : a
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
24 pi2 : b
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
25
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
26 open _/\_
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
27
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
28 _-_ : ℕ → ℕ → ℕ
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
29 x - zero = x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
30 zero - _ = zero
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
31 (suc x) - (suc y) = x - y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
32
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
33 +zero : { y : ℕ } → y + zero ≡ y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
34 +zero {zero} = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
35 +zero {suc y} = cong ( λ x → suc x ) ( +zero {y} )
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
36
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
37
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
38 +-sym : { x y : ℕ } → x + y ≡ y + x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
39 +-sym {zero} {zero} = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
40 +-sym {zero} {suc y} = let open ≡-Reasoning in
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
41 begin
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
42 zero + suc y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
43 ≡⟨⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
44 suc y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
45 ≡⟨ sym +zero ⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
46 suc y + zero
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
47
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
48 +-sym {suc x} {zero} = let open ≡-Reasoning in
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
49 begin
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
50 suc x + zero
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
51 ≡⟨ +zero ⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
52 suc x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
53 ≡⟨⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
54 zero + suc x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
55
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
56 +-sym {suc x} {suc y} = cong ( λ z → suc z ) ( let open ≡-Reasoning in
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
57 begin
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
58 x + suc y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
59 ≡⟨ +-sym {x} {suc y} ⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
60 suc (y + x)
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
61 ≡⟨ cong ( λ z → suc z ) (+-sym {y} {x}) ⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
62 suc (x + y)
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
63 ≡⟨ sym ( +-sym {y} {suc x}) ⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
64 y + suc x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
65 ∎ )
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
66
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
67 minus-plus : { x y : ℕ } → (suc x - 1) + (y + 1) ≡ suc x + y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
68 minus-plus {zero} {y} = +-sym {y} {1}
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
69 minus-plus {suc x} {y} = cong ( λ z → suc z ) (minus-plus {x} {y})
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
70
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
71 +1≡suc : { x : ℕ } → x + 1 ≡ suc x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
72 +1≡suc {zero} = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
73 +1≡suc {suc x} = cong ( λ z → suc z ) ( +1≡suc {x} )
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
74
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
75 lt : ℕ → ℕ → Bool
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
76 lt x y with (suc x ) ≤? y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
77 lt x y | yes p = true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
78 lt x y | no ¬p = false
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
79
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
80 predℕ : {n : ℕ } → lt 0 n ≡ true → ℕ
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
81 predℕ {zero} ()
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
82 predℕ {suc n} refl = n
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
83
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
84 predℕ+1=n : {n : ℕ } → (less : lt 0 n ≡ true ) → (predℕ less) + 1 ≡ n
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
85 predℕ+1=n {zero} ()
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
86 predℕ+1=n {suc n} refl = +1≡suc
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
87
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
88 suc-predℕ=n : {n : ℕ } → (less : lt 0 n ≡ true ) → suc (predℕ less) ≡ n
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
89 suc-predℕ=n {zero} ()
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
90 suc-predℕ=n {suc n} refl = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
91
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
92 Equal : ℕ → ℕ → Bool
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
93 Equal x y with x ≟ y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
94 Equal x y | yes p = true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
95 Equal x y | no ¬p = false
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
96
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
97 _⇒_ : Bool → Bool → Bool
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
98 false ⇒ _ = true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
99 true ⇒ true = true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
100 true ⇒ false = false
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
101
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
102 ⇒t : {x : Bool} → x ⇒ true ≡ true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
103 ⇒t {x} with x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
104 ⇒t {x} | false = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
105 ⇒t {x} | true = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
106
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
107 f⇒ : {x : Bool} → false ⇒ x ≡ true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
108 f⇒ {x} with x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
109 f⇒ {x} | false = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
110 f⇒ {x} | true = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
111
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
112 ∧-pi1 : { x y : Bool } → x ∧ y ≡ true → x ≡ true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
113 ∧-pi1 {x} {y} eq with x | y | eq
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
114 ∧-pi1 {x} {y} eq | false | b | ()
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
115 ∧-pi1 {x} {y} eq | true | false | ()
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
116 ∧-pi1 {x} {y} eq | true | true | refl = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
117
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
118 ∧-pi2 : { x y : Bool } → x ∧ y ≡ true → y ≡ true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
119 ∧-pi2 {x} {y} eq with x | y | eq
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
120 ∧-pi2 {x} {y} eq | false | b | ()
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
121 ∧-pi2 {x} {y} eq | true | false | ()
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
122 ∧-pi2 {x} {y} eq | true | true | refl = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
123
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
124 ∧true : { x : Bool } → x ∧ true ≡ x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
125 ∧true {x} with x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
126 ∧true {x} | false = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
127 ∧true {x} | true = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
128
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
129 true∧ : { x : Bool } → true ∧ x ≡ x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
130 true∧ {x} with x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
131 true∧ {x} | false = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
132 true∧ {x} | true = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
133 bool-case : ( x : Bool ) { p : Set } → ( x ≡ true → p ) → ( x ≡ false → p ) → p
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
134 bool-case x T F with x
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
135 bool-case x T F | false = F refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
136 bool-case x T F | true = T refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
137
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
138 impl⇒ : {x y : Bool} → (x ≡ true → y ≡ true ) → x ⇒ y ≡ true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
139 impl⇒ {x} {y} p = bool-case x (λ x=t → let open ≡-Reasoning in
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
140 begin
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
141 x ⇒ y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
142 ≡⟨ cong ( λ z → x ⇒ z ) (p x=t ) ⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
143 x ⇒ true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
144 ≡⟨ ⇒t ⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
145 true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
146
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
147 ) ( λ x=f → let open ≡-Reasoning in
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
148 begin
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
149 x ⇒ y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
150 ≡⟨ cong ( λ z → z ⇒ y ) x=f ⟩
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
151 true
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
152
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
153 )
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
154
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
155 Equal→≡ : { x y : ℕ } → Equal x y ≡ true → x ≡ y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
156 Equal→≡ {x} {y} eq with x ≟ y
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
157 Equal→≡ {x} {y} refl | yes refl = refl
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
158 Equal→≡ {x} {y} () | no ¬p
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
159
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
160 Equal+1 : { x y : ℕ } → Equal x y ≡ Equal (suc x) (suc y)
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
161 Equal+1 {x} {y} with x ≟ y
27
a39a82820742 add whileTestCondition
ryokka
parents: 21
diff changeset
162 Equal+1 {x} {.x} | yes refl = {!!}
a39a82820742 add whileTestCondition
ryokka
parents: 21
diff changeset
163 Equal+1 {x} {y} | no ¬p = {!!}
12
247ce3e67b5f add utilitites
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
diff changeset
164
15
8d546766a9a8 Prim variable version done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 12
diff changeset
165 open import Data.Empty
8d546766a9a8 Prim variable version done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 12
diff changeset
166
8d546766a9a8 Prim variable version done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 12
diff changeset
167 ≡→Equal : { x y : ℕ } → x ≡ y → Equal x y ≡ true
8d546766a9a8 Prim variable version done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 12
diff changeset
168 ≡→Equal {x} {.x} refl with x ≟ x
8d546766a9a8 Prim variable version done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 12
diff changeset
169 ≡→Equal {x} {.x} refl | yes refl = refl
8d546766a9a8 Prim variable version done
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents: 12
diff changeset
170 ≡→Equal {x} {.x} refl | no ¬p = ⊥-elim ( ¬p refl )