Mercurial > hg > Members > ryokka > HoareLogic
annotate whileTestGears.agda @ 45:de8449321356
close
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Sun, 15 Dec 2019 21:42:28 +0900 |
parents | 5a3c9d087c7c |
children |
rev | line source |
---|---|
4 | 1 module whileTestGears where |
2 | |
3 open import Function | |
4 open import Data.Nat | |
34 | 5 open import Data.Bool hiding ( _≟_ ; _≤?_ ; _≤_ ; _<_) |
4 | 6 open import Level renaming ( suc to succ ; zero to Zero ) |
7 open import Relation.Nullary using (¬_; Dec; yes; no) | |
8 open import Relation.Binary.PropositionalEquality | |
9 | |
10 | 10 open import utilities |
11 open _/\_ | |
4 | 12 |
44 | 13 record Env (Cxt : Set) : Set (succ Zero) where |
6 | 14 field |
15 varn : ℕ | |
16 vari : ℕ | |
44 | 17 cx : Cxt |
18 open Env | |
6 | 19 |
44 | 20 whileTest : {l : Level} {t : Set l} {Cxt : Set} (c : Cxt ) → (c10 : ℕ) → (Code : Env Cxt → t) → t |
21 whileTest c c10 next = next (record {varn = c10 ; vari = 0 ; cx = c } ) | |
4 | 22 |
23 {-# TERMINATING #-} | |
44 | 24 whileLoop : {l : Level} {t : Set l} {Cxt : Set} {c : Cxt } → Env Cxt → (Code : Env Cxt → t) → t |
4 | 25 whileLoop env next with lt 0 (varn env) |
26 whileLoop env next | false = next env | |
27 whileLoop env next | true = | |
42 | 28 whileLoop (record env {varn = (varn env) - 1 ; vari = (vari env) + 1}) next |
4 | 29 |
44 | 30 test1 : {Cxt : Set } → (c : Cxt) → Env Cxt |
42 | 31 test1 c = whileTest c 10 (λ env → whileLoop env (λ env1 → env1 )) |
4 | 32 |
33 | |
44 | 34 proof1 : {Cxt : Set } (c : Cxt ) → whileTest c 10 (λ env → whileLoop env (λ e → (vari e) ≡ 10 )) |
42 | 35 proof1 c = refl |
4 | 36 |
16 | 37 -- ↓PostCondition |
44 | 38 whileTest' : {l : Level} {t : Set l} {Cxt : Set} → {c : Cxt} → {c10 : ℕ } → (Code : (env : Env Cxt ) → ((vari env) ≡ 0) /\ ((varn env) ≡ c10) → t) → t |
39 whileTest' {_} {_} {Cxt} {c} {c10} next = next env proof2 | |
4 | 40 where |
44 | 41 env : Env Cxt |
42 | 42 env = record {vari = 0 ; varn = c10 ; cx = c } |
16 | 43 proof2 : ((vari env) ≡ 0) /\ ((varn env) ≡ c10) -- PostCondition |
4 | 44 proof2 = record {pi1 = refl ; pi2 = refl} |
11 | 45 |
46 open import Data.Empty | |
47 open import Data.Nat.Properties | |
48 | |
49 | |
16 | 50 {-# TERMINATING #-} -- ↓PreCondition(Invaliant) |
44 | 51 whileLoop' : {l : Level} {t : Set l} {Cxt : Set} {c : Cxt } → (env : Env Cxt ) → {c10 : ℕ } → ((varn env) + (vari env) ≡ c10) → (Code : Env Cxt → t) → t |
9 | 52 whileLoop' env proof next with ( suc zero ≤? (varn env) ) |
53 whileLoop' env proof next | no p = next env | |
14 | 54 whileLoop' env {c10} proof next | yes p = whileLoop' env1 (proof3 p ) next |
4 | 55 where |
42 | 56 env1 = record env {varn = (varn env) - 1 ; vari = (vari env) + 1} |
11 | 57 1<0 : 1 ≤ zero → ⊥ |
58 1<0 () | |
14 | 59 proof3 : (suc zero ≤ (varn env)) → varn env1 + vari env1 ≡ c10 |
42 | 60 proof3 = {!!} |
61 -- proof3 (s≤s lt) with varn env | |
62 -- proof3 (s≤s z≤n) | zero = ⊥-elim (1<0 p) | |
63 -- proof3 (s≤s (z≤n {n'}) ) | suc n = let open ≡-Reasoning in | |
64 -- begin | |
65 -- n' + (vari env + 1) | |
66 -- ≡⟨ cong ( λ z → n' + z ) ( +-sym {vari env} {1} ) ⟩ | |
67 -- n' + (1 + vari env ) | |
68 -- ≡⟨ sym ( +-assoc (n') 1 (vari env) ) ⟩ | |
69 -- (n' + 1) + vari env | |
70 -- ≡⟨ cong ( λ z → z + vari env ) +1≡suc ⟩ | |
71 -- (suc n' ) + vari env | |
72 -- ≡⟨⟩ | |
73 -- varn env + vari env | |
74 -- ≡⟨ proof ⟩ | |
75 -- c10 | |
76 -- ∎ | |
6 | 77 |
16 | 78 -- Condition to Invaliant |
44 | 79 conversion1 : {l : Level} {t : Set l } {Cxt : Set} {c : Cxt } → (env : Env Cxt ) → {c10 : ℕ } → ((vari env) ≡ 0) /\ ((varn env) ≡ c10) |
80 → (Code : (env1 : Env Cxt ) → (varn env1 + vari env1 ≡ c10) → t) → t | |
14 | 81 conversion1 env {c10} p1 next = next env proof4 |
6 | 82 where |
14 | 83 proof4 : varn env + vari env ≡ c10 |
6 | 84 proof4 = let open ≡-Reasoning in |
85 begin | |
86 varn env + vari env | |
87 ≡⟨ cong ( λ n → n + vari env ) (pi2 p1 ) ⟩ | |
14 | 88 c10 + vari env |
89 ≡⟨ cong ( λ n → c10 + n ) (pi1 p1 ) ⟩ | |
90 c10 + 0 | |
91 ≡⟨ +-sym {c10} {0} ⟩ | |
92 c10 | |
6 | 93 ∎ |
4 | 94 |
6 | 95 |
44 | 96 proofGears : {c10 : ℕ } → { Cxt : Set } → (c : Cxt ) → Set |
97 proofGears {c10} c = whileTest' {_} {_} {_} {c} {c10} (λ n p1 → conversion1 n p1 (λ n1 p2 → whileLoop' n1 p2 (λ n2 → ( vari n2 ≡ c10 )))) | |
9 | 98 |
44 | 99 record Context (e : Set ) : Set (succ Zero) |
100 | |
101 data whileTestState {Cxt : Set } (c10 : ℕ ) (env : Env Cxt ) : Set where | |
33 | 102 error : whileTestState c10 env |
103 state1 : ((vari env) ≡ 0) /\ ((varn env) ≡ c10) → whileTestState c10 env | |
104 state2 : (varn env + vari env ≡ c10) → whileTestState c10 env | |
105 finstate : ((vari env) ≡ c10 ) → whileTestState c10 env | |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
106 |
41 | 107 -- |
44 | 108 -- openended Env Cxt c <=> Context |
41 | 109 -- |
110 | |
44 | 111 record Context e where |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
112 field |
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
113 c10 : ℕ |
44 | 114 whileDG : Env e |
33 | 115 whileCond : whileTestState c10 whileDG |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
116 |
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
117 open Context |
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
118 |
41 | 119 open import Relation.Nullary |
120 open import Relation.Binary | |
121 | |
122 -- | |
123 -- transparency of condition | |
124 -- | |
125 | |
44 | 126 whileCondP : Env {!!} → Set |
41 | 127 whileCondP env = varn env > 0 |
128 | |
129 whileDec : (cxt : Context) → Dec (whileCondP (whileDG cxt)) | |
130 whileDec cxt = {!!} | |
131 | |
33 | 132 whileTestContext : {l : Level} {t : Set l} → Context → (Code : Context → t) → t |
133 whileTestContext cxt next = next (record cxt { whileDG = record (whileDG cxt) {varn = c10 cxt ; vari = 0} ; whileCond = {!!} } ) | |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
134 |
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
135 {-# TERMINATING #-} |
33 | 136 whileLoopContext : {l : Level} {t : Set l} → Context → (Code : Context → t) → t |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
137 whileLoopContext cxt next with lt 0 (varn (whileDG cxt) ) |
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
138 whileLoopContext cxt next | false = next cxt |
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
139 whileLoopContext cxt next | true = |
42 | 140 whileLoopContext (record cxt { whileDG = record (whileDG cxt) {varn = (varn (whileDG cxt)) - 1 ; vari = (vari (whileDG cxt)) + 1} ; whileCond = state2 (proof cxt) } ) next |
38 | 141 where |
142 proof : (cxt : Context) → (varn (whileDG cxt) - 1) + (vari (whileDG cxt) + 1) ≡ c10 cxt | |
143 proof cxt = {!!} | |
33 | 144 |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
145 {-# TERMINATING #-} |
44 | 146 whileLoopStep : {l : Level} {t : Set l} → Env {!!} → (next : (e : Env {!!} ) → t) (exit : (e : Env {!!} ) → t) → t |
33 | 147 whileLoopStep env next exit with <-cmp 0 (varn env) |
41 | 148 whileLoopStep env next exit | tri≈ _ eq _ = exit env |
42 | 149 whileLoopStep env next exit | tri< gt ¬eq _ = next record env {varn = (varn env) - 1 ; vari = (vari env) + 1} |
39 | 150 whileLoopStep env next exit | tri> _ _ c = ⊥-elim (m<n⇒n≢0 {varn env} {0} c refl) |
27 | 151 |
40 | 152 whileTestProof : {l : Level} {t : Set l} → Context → (Code : (cxt : Context ) → ¬ (vari (whileDG cxt) ≡ varn (whileDG cxt) ) → t) → t |
153 whileTestProof cxt next = next record cxt { whileDG = out ; whileCond = init } i!=n where | |
44 | 154 out : Env {!!} |
42 | 155 out = whileTest {!!} (c10 cxt) ( λ e → e ) |
33 | 156 init : whileTestState (c10 cxt) out |
157 init = state1 record { pi1 = refl ; pi2 = refl } | |
40 | 158 i!=n : ¬ vari out ≡ varn out |
159 i!=n eq = {!!} | |
27 | 160 |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
161 {-# TERMINATING #-} |
41 | 162 whileLoopProof : {l : Level} {t : Set l} → (cxt : Context ) → whileCondP (whileDG cxt) |
163 → (continue : (cxt : Context ) → whileCondP (whileDG cxt) → t) (exit : Context → ¬ whileCondP (whileDG cxt) → t) → t | |
40 | 164 whileLoopProof cxt i!=n next exit = whileLoopStep (whileDG cxt) |
41 | 165 ( λ env → next record cxt { whileDG = env ; whileCond = {!!} } {!!} ) |
166 ( λ env → exit record cxt { whileDG = env ; whileCond = exitCond env {!!} } {!!} ) where | |
44 | 167 proof5 : (e : Env {!!} ) → varn e + vari e ≡ c10 cxt → 0 ≡ varn e → vari e ≡ c10 cxt |
34 | 168 proof5 record { varn = .0 ; vari = vari } refl refl = refl |
44 | 169 exitCond : (e : Env {!!} ) → 0 ≡ varn e → whileTestState (c10 cxt) e |
33 | 170 exitCond nenv eq1 with whileCond cxt | inspect whileDG cxt |
171 ... | state2 cond | record { eq = eq2 } = finstate ( proof5 nenv {!!} eq1 ) | |
172 ... | _ | _ = error | |
27 | 173 |
40 | 174 whileConvProof : {l : Level} {t : Set l} → (cxt : Context ) → ¬ (vari (whileDG cxt) ≡ varn (whileDG cxt)) |
175 → ((cxt : Context ) → ¬ (vari (whileDG cxt) ≡ varn (whileDG cxt)) → t) → t | |
176 whileConvProof cxt i!=n next = next record cxt { whileCond = postCond } i!=n where | |
44 | 177 proof4 : (e : Env {!!} ) → (vari e ≡ 0) /\ (varn e ≡ c10 cxt) → varn e + vari e ≡ c10 cxt |
32
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
178 proof4 env p1 = let open ≡-Reasoning in |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
179 begin |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
180 varn env + vari env |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
181 ≡⟨ cong ( λ n → n + vari env ) (pi2 p1 ) ⟩ |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
182 c10 cxt + vari env |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
183 ≡⟨ cong ( λ n → c10 cxt + n ) (pi1 p1 ) ⟩ |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
184 c10 cxt + 0 |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
185 ≡⟨ +-sym {c10 cxt} {0} ⟩ |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
186 c10 cxt |
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
187 ∎ |
33 | 188 postCond : whileTestState (c10 cxt) (whileDG cxt) |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
189 postCond with whileCond cxt |
33 | 190 ... | state1 cond = state2 (proof4 (whileDG cxt) cond ) |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
191 ... | _ = error |
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
192 |
41 | 193 {-# TERMINATING #-} |
194 loop : {l : Level} {t : Set l} → (cxt : Context ) → whileCondP (whileDG cxt) | |
195 → (exit : (cxt : Context ) → ¬ whileCondP (whileDG cxt) → t) → t | |
196 loop cxt i!=n exit = whileLoopProof cxt i!=n (λ cxt → loop cxt {!!} {!!} exit ) {!!} | |
30
dd66b94bf365
loop causes agda inifinite loop
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
29
diff
changeset
|
197 |
35 | 198 |
36 | 199 {-# TERMINATING #-} |
37 | 200 loopProof : {l : Level} {t : Set l} {P Inv : Context → Set } → (c : Context) → (if : (c : Context) → Dec (P c)) |
41 | 201 → ( (cont : (c2 : Context) → (P c2) → t) → (exit : (c2 : Context) → ¬ (P c2) → t) → t ) |
202 → t | |
203 loopProof {l} {t} {P} {Inv} cxt if f = {!!} | |
36 | 204 where |
41 | 205 lem : (c : Context) → t |
206 lem c with if c | |
207 lem c | no ¬p = {!!} -- f c (λ c1 _ → lem c1 ) | |
208 lem c | yes p = {!!} | |
36 | 209 |
42 | 210 whileLoopProof' : {l : Level} {t : Set l} |
211 → (cxt : Context ) → ( continue : (cxt : Context ) → whileCondP (whileDG cxt) → t) (exit : Context → ¬ whileCondP (whileDG cxt) → t) → t | |
212 whileLoopProof' = {!!} | |
35 | 213 |
41 | 214 proofWhileGear : (c : ℕ) (cxt : Context) → whileTestProof (record cxt { c10 = c ; whileCond = error }) |
215 $ λ cxt i!=n → whileConvProof cxt i!=n | |
42 | 216 $ λ cxt i+n=c10 → loopProof cxt whileDec {!!} -- whileLoopProof' |
217 $ λ cxt _ → {!!} | |
32
bf7c7bd69e35
conversion. loop needs cases
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
31
diff
changeset
|
218 proofWhileGear c cxt = {!!} |