3
|
1 module whileTestGears where
|
|
2
|
|
3 open import Function
|
|
4 open import Data.Nat
|
|
5 open import Data.Bool hiding ( _@$\stackrel{?}{=}$@_ ; _@$\leq$@?_ ; _@$\leq$@_ ; _<_)
|
|
6 open import Data.Product
|
|
7 open import Level renaming ( suc to succ ; zero to Zero )
|
|
8 open import Relation.Nullary using (@$\neg$@_; Dec; yes; no)
|
|
9 open import Relation.Binary.PropositionalEquality
|
|
10 open import Agda.Builtin.Unit
|
|
11
|
|
12 open import utilities
|
|
13 open _@$\wedge$@_
|
|
14
|
|
15 -- original codeGear (with non terminatinng )
|
|
16
|
|
17 record Env : Set (succ Zero) where
|
|
18 field
|
|
19 varn : @$\mathbb{N}$@
|
|
20 vari : @$\mathbb{N}$@
|
|
21 open Env
|
|
22
|
|
23 whileTest : {l : Level} {t : Set l} @$\rightarrow$@ (c10 : @$\mathbb{N}$@) @$\rightarrow$@ (Code : Env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
24 whileTest c10 next = next (record {varn = c10 ; vari = 0 } )
|
|
25
|
|
26 {-@$\#$@ TERMINATING @$\#$@-}
|
|
27 whileLoop : {l : Level} {t : Set l} @$\rightarrow$@ Env @$\rightarrow$@ (Code : Env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
28 whileLoop env next with lt 0 (varn env)
|
|
29 whileLoop env next | false = next env
|
|
30 whileLoop env next | true =
|
|
31 whileLoop (record env {varn = (varn env) - 1 ; vari = (vari env) + 1}) next
|
|
32
|
|
33 test1 : Env
|
|
34 test1 = whileTest 10 (@$\lambda$@ env @$\rightarrow$@ whileLoop env (@$\lambda$@ env1 @$\rightarrow$@ env1 ))
|
|
35
|
|
36 proof1 : whileTest 10 (@$\lambda$@ env @$\rightarrow$@ whileLoop env (@$\lambda$@ e @$\rightarrow$@ (vari e) @$\equiv$@ 10 ))
|
|
37 proof1 = refl
|
|
38
|
|
39 -- codeGear with pre-condtion and post-condition
|
|
40 --
|
|
41 -- ↓PostCondition
|
|
42 whileTest' : {l : Level} {t : Set l} @$\rightarrow$@ {c10 : @$\mathbb{N}$@ } @$\rightarrow$@ (Code : (env : Env ) @$\rightarrow$@ ((vari env) @$\equiv$@ 0) @$\wedge$@ ((varn env) @$\equiv$@ c10) @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
43 whileTest' {_} {_} {c10} next = next env proof2
|
|
44 where
|
|
45 env : Env
|
|
46 env = record {vari = 0 ; varn = c10 }
|
|
47 proof2 : ((vari env) @$\equiv$@ 0) @$\wedge$@ ((varn env) @$\equiv$@ c10) -- PostCondition
|
|
48 proof2 = record {pi1 = refl ; pi2 = refl}
|
|
49
|
|
50
|
|
51 open import Data.Empty
|
|
52 open import Data.Nat.Properties
|
|
53
|
|
54
|
|
55 {-@$\#$@ TERMINATING @$\#$@-} -- ↓PreCondition(Invaliant)
|
|
56 whileLoop' : {l : Level} {t : Set l} @$\rightarrow$@ (env : Env ) @$\rightarrow$@ {c10 : @$\mathbb{N}$@ } @$\rightarrow$@ ((varn env) + (vari env) @$\equiv$@ c10) @$\rightarrow$@ (Code : Env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
57 whileLoop' env proof next with ( suc zero @$\leq$@? (varn env) )
|
|
58 whileLoop' env proof next | no p = next env
|
|
59 whileLoop' env {c10} proof next | yes p = whileLoop' env1 (proof3 p ) next
|
|
60 where
|
|
61 env1 = record env {varn = (varn env) - 1 ; vari = (vari env) + 1}
|
|
62 1<0 : 1 @$\leq$@ zero @$\rightarrow$@ @$\bot$@
|
|
63 1<0 ()
|
|
64 proof3 : (suc zero @$\leq$@ (varn env)) @$\rightarrow$@ varn env1 + vari env1 @$\equiv$@ c10
|
|
65 proof3 (s@$\leq$@s lt) with varn env
|
|
66 proof3 (s@$\leq$@s z@$\leq$@n) | zero = @$\bot$@-elim (1<0 p)
|
|
67 proof3 (s@$\leq$@s (z@$\leq$@n {n'}) ) | suc n = let open @$\equiv$@-Reasoning in
|
|
68 begin
|
|
69 n' + (vari env + 1)
|
|
70 @$\equiv$@@$\langle$@ cong ( @$\lambda$@ z @$\rightarrow$@ n' + z ) ( +-sym {vari env} {1} ) @$\rangle$@
|
|
71 n' + (1 + vari env )
|
|
72 @$\equiv$@@$\langle$@ sym ( +-assoc (n') 1 (vari env) ) @$\rangle$@
|
|
73 (n' + 1) + vari env
|
|
74 @$\equiv$@@$\langle$@ cong ( @$\lambda$@ z @$\rightarrow$@ z + vari env ) +1@$\equiv$@suc @$\rangle$@
|
|
75 (suc n' ) + vari env
|
|
76 @$\equiv$@@$\langle$@@$\rangle$@
|
|
77 varn env + vari env
|
|
78 @$\equiv$@@$\langle$@ proof @$\rangle$@
|
|
79 c10
|
|
80 @$\blacksquare$@
|
|
81
|
|
82 -- Condition to Invariant
|
|
83 conversion1 : {l : Level} {t : Set l } @$\rightarrow$@ (env : Env ) @$\rightarrow$@ {c10 : @$\mathbb{N}$@ } @$\rightarrow$@ ((vari env) @$\equiv$@ 0) @$\wedge$@ ((varn env) @$\equiv$@ c10)
|
|
84 @$\rightarrow$@ (Code : (env1 : Env ) @$\rightarrow$@ (varn env1 + vari env1 @$\equiv$@ c10) @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
85 conversion1 env {c10} p1 next = next env proof4
|
|
86 where
|
|
87 proof4 : varn env + vari env @$\equiv$@ c10
|
|
88 proof4 = let open @$\equiv$@-Reasoning in
|
|
89 begin
|
|
90 varn env + vari env
|
|
91 @$\equiv$@@$\langle$@ cong ( @$\lambda$@ n @$\rightarrow$@ n + vari env ) (pi2 p1 ) @$\rangle$@
|
|
92 c10 + vari env
|
|
93 @$\equiv$@@$\langle$@ cong ( @$\lambda$@ n @$\rightarrow$@ c10 + n ) (pi1 p1 ) @$\rangle$@
|
|
94 c10 + 0
|
|
95 @$\equiv$@@$\langle$@ +-sym {c10} {0} @$\rangle$@
|
|
96 c10
|
|
97 @$\blacksquare$@
|
|
98
|
|
99 -- all proofs are connected
|
|
100 proofGears : {c10 : @$\mathbb{N}$@ } @$\rightarrow$@ Set
|
|
101 proofGears {c10} = whileTest' {_} {_} {c10} (@$\lambda$@ n p1 @$\rightarrow$@ conversion1 n p1 (@$\lambda$@ n1 p2 @$\rightarrow$@ whileLoop' n1 p2 (@$\lambda$@ n2 @$\rightarrow$@ ( vari n2 @$\equiv$@ c10 ))))
|
|
102
|
|
103 --
|
|
104 -- codeGear with loop step and closed environment
|
|
105 --
|
|
106
|
|
107 open import Relation.Binary
|
|
108
|
|
109 record Envc : Set (succ Zero) where
|
|
110 field
|
|
111 c10 : @$\mathbb{N}$@
|
|
112 varn : @$\mathbb{N}$@
|
|
113 vari : @$\mathbb{N}$@
|
|
114 open Envc
|
|
115
|
|
116 whileTestP : {l : Level} {t : Set l} @$\rightarrow$@ (c10 : @$\mathbb{N}$@) @$\rightarrow$@ (Code : Envc @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
117 whileTestP c10 next = next (record {varn = c10 ; vari = 0 ; c10 = c10 } )
|
|
118
|
|
119 whileLoopP : {l : Level} {t : Set l} @$\rightarrow$@ Envc @$\rightarrow$@ (next : Envc @$\rightarrow$@ t) @$\rightarrow$@ (exit : Envc @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
120 whileLoopP env next exit with <-cmp 0 (varn env)
|
|
121 whileLoopP env next exit | tri@$\thickapprox$@ @$\neg$@a b @$\neg$@c = exit env
|
|
122 whileLoopP env next exit | tri< a @$\neg$@b @$\neg$@c =
|
|
123 next (record env {varn = (varn env) - 1 ; vari = (vari env) + 1 })
|
|
124
|
|
125 -- equivalent of whileLoopP but it looks like an induction on varn
|
|
126 whileLoopP' : {l : Level} {t : Set l} @$\rightarrow$@ (n : @$\mathbb{N}$@) @$\rightarrow$@ (env : Envc) @$\rightarrow$@ (n @$\equiv$@ varn env) @$\rightarrow$@ (next : Envc @$\rightarrow$@ t) @$\rightarrow$@ (exit : Envc @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
127 whileLoopP' zero env refl _ exit = exit env
|
|
128 whileLoopP' (suc n) env refl next _ = next (record {c10 = (c10 env) ; varn = varn env ; vari = suc (vari env) })
|
|
129
|
|
130 -- normal loop without termination
|
|
131 {-@$\#$@ TERMINATING @$\#$@-}
|
|
132 loopP : {l : Level} {t : Set l} @$\rightarrow$@ Envc @$\rightarrow$@ (exit : Envc @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
133 loopP env exit = whileLoopP env (@$\lambda$@ env @$\rightarrow$@ loopP env exit ) exit
|
|
134
|
|
135 whileTestPCall : (c10 : @$\mathbb{N}$@ ) @$\rightarrow$@ Envc
|
|
136 whileTestPCall c10 = whileTestP {_} {_} c10 (@$\lambda$@ env @$\rightarrow$@ loopP env (@$\lambda$@ env @$\rightarrow$@ env))
|
|
137
|
|
138 --
|
|
139 -- codeGears with states of condition
|
|
140 --
|
|
141 data whileTestState : Set where
|
|
142 s1 : whileTestState
|
|
143 s2 : whileTestState
|
|
144 sf : whileTestState
|
|
145
|
|
146 whileTestStateP : whileTestState @$\rightarrow$@ Envc @$\rightarrow$@ Set
|
|
147 whileTestStateP s1 env = (vari env @$\equiv$@ 0) @$\wedge$@ (varn env @$\equiv$@ c10 env)
|
|
148 whileTestStateP s2 env = (varn env + vari env @$\equiv$@ c10 env)
|
|
149 whileTestStateP sf env = (vari env @$\equiv$@ c10 env)
|
|
150
|
|
151 whileTestPwP : {l : Level} {t : Set l} @$\rightarrow$@ (c10 : @$\mathbb{N}$@) @$\rightarrow$@ ((env : Envc ) @$\rightarrow$@ whileTestStateP s1 env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
152 whileTestPwP c10 next = next env record { pi1 = refl ; pi2 = refl } where
|
|
153 env : Envc
|
|
154 env = whileTestP c10 ( @$\lambda$@ env @$\rightarrow$@ env )
|
|
155
|
|
156 whileLoopPwP : {l : Level} {t : Set l} @$\rightarrow$@ (env : Envc ) @$\rightarrow$@ whileTestStateP s2 env
|
|
157 @$\rightarrow$@ (next : (env : Envc ) @$\rightarrow$@ whileTestStateP s2 env @$\rightarrow$@ t)
|
|
158 @$\rightarrow$@ (exit : (env : Envc ) @$\rightarrow$@ whileTestStateP sf env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
159 whileLoopPwP env s next exit with <-cmp 0 (varn env)
|
|
160 whileLoopPwP env s next exit | tri@$\thickapprox$@ @$\neg$@a b @$\neg$@c = exit env (lem (sym b) s)
|
|
161 where
|
|
162 lem : (varn env @$\equiv$@ 0) @$\rightarrow$@ (varn env + vari env @$\equiv$@ c10 env) @$\rightarrow$@ vari env @$\equiv$@ c10 env
|
|
163 lem refl refl = refl
|
|
164 whileLoopPwP env s next exit | tri< a @$\neg$@b @$\neg$@c = next (record env {varn = (varn env) - 1 ; vari = (vari env) + 1 }) (proof5 a)
|
|
165 where
|
|
166 1<0 : 1 @$\leq$@ zero @$\rightarrow$@ @$\bot$@
|
|
167 1<0 ()
|
|
168 proof5 : (suc zero @$\leq$@ (varn env)) @$\rightarrow$@ (varn env - 1) + (vari env + 1) @$\equiv$@ c10 env
|
|
169 proof5 (s@$\leq$@s lt) with varn env
|
|
170 proof5 (s@$\leq$@s z@$\leq$@n) | zero = @$\bot$@-elim (1<0 a)
|
|
171 proof5 (s@$\leq$@s (z@$\leq$@n {n'}) ) | suc n = let open @$\equiv$@-Reasoning in
|
|
172 begin
|
|
173 n' + (vari env + 1)
|
|
174 @$\equiv$@@$\langle$@ cong ( @$\lambda$@ z @$\rightarrow$@ n' + z ) ( +-sym {vari env} {1} ) @$\rangle$@
|
|
175 n' + (1 + vari env )
|
|
176 @$\equiv$@@$\langle$@ sym ( +-assoc (n') 1 (vari env) ) @$\rangle$@
|
|
177 (n' + 1) + vari env
|
|
178 @$\equiv$@@$\langle$@ cong ( @$\lambda$@ z @$\rightarrow$@ z + vari env ) +1@$\equiv$@suc @$\rangle$@
|
|
179 (suc n' ) + vari env
|
|
180 @$\equiv$@@$\langle$@@$\rangle$@
|
|
181 varn env + vari env
|
|
182 @$\equiv$@@$\langle$@ s @$\rangle$@
|
|
183 c10 env
|
|
184 @$\blacksquare$@
|
|
185
|
|
186
|
|
187 whileLoopPwP' : {l : Level} {t : Set l} @$\rightarrow$@ (n : @$\mathbb{N}$@) @$\rightarrow$@ (env : Envc ) @$\rightarrow$@ (n @$\equiv$@ varn env) @$\rightarrow$@ whileTestStateP s2 env
|
|
188 @$\rightarrow$@ (next : (env : Envc ) @$\rightarrow$@ (pred n @$\equiv$@ varn env) @$\rightarrow$@ whileTestStateP s2 env @$\rightarrow$@ t)
|
|
189 @$\rightarrow$@ (exit : (env : Envc ) @$\rightarrow$@ whileTestStateP sf env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
190 whileLoopPwP' zero env refl refl next exit = exit env refl
|
|
191 whileLoopPwP' (suc n) env refl refl next exit = next (record env {varn = pred (varn env) ; vari = suc (vari env) }) refl (+-suc n (vari env))
|
|
192
|
|
193
|
|
194 {-@$\#$@ TERMINATING @$\#$@-}
|
|
195 loopPwP : {l : Level} {t : Set l} @$\rightarrow$@ (env : Envc ) @$\rightarrow$@ whileTestStateP s2 env @$\rightarrow$@ (exit : (env : Envc ) @$\rightarrow$@ whileTestStateP sf env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
196 loopPwP env s exit = whileLoopPwP env s (@$\lambda$@ env s @$\rightarrow$@ loopPwP env s exit ) exit
|
|
197
|
|
198
|
|
199 loopPwP' : {l : Level} {t : Set l} @$\rightarrow$@ (n : @$\mathbb{N}$@) @$\rightarrow$@ (env : Envc ) @$\rightarrow$@ (n @$\equiv$@ varn env) @$\rightarrow$@ whileTestStateP s2 env @$\rightarrow$@ (exit : (env : Envc ) @$\rightarrow$@ whileTestStateP sf env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
200 loopPwP' zero env refl refl exit = exit env refl
|
|
201 loopPwP' (suc n) env refl refl exit = whileLoopPwP' (suc n) env refl refl (@$\lambda$@ env x y @$\rightarrow$@ loopPwP' n env x y exit) exit
|
|
202
|
|
203
|
|
204 loopHelper : (n : @$\mathbb{N}$@) @$\rightarrow$@ (env : Envc ) @$\rightarrow$@ (eq : varn env @$\equiv$@ n) @$\rightarrow$@ (seq : whileTestStateP s2 env) @$\rightarrow$@ loopPwP' n env (sym eq) seq @$\lambda$@ env@$\_{1}$@ x @$\rightarrow$@ (vari env@$\_{1}$@ @$\equiv$@ c10 env@$\_{1}$@)
|
|
205 loopHelper zero env eq refl rewrite eq = refl
|
|
206 loopHelper (suc n) env eq refl rewrite eq = loopHelper n (record { c10 = suc (n + vari env) ; varn = n ; vari = suc (vari env) }) refl (+-suc n (vari env))
|
|
207
|
|
208
|
|
209 -- all codtions are correctly connected and required condtion is proved in the continuation
|
|
210 -- use required condition as t in (env @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
211 --
|
|
212 whileTestPCallwP : (c : @$\mathbb{N}$@ ) @$\rightarrow$@ Set
|
|
213 whileTestPCallwP c = whileTestPwP {_} {_} c ( @$\lambda$@ env s @$\rightarrow$@ loopPwP env (conv env s) ( @$\lambda$@ env s @$\rightarrow$@ vari env @$\equiv$@ c10 env ) ) where
|
|
214 conv : (env : Envc ) @$\rightarrow$@ (vari env @$\equiv$@ 0) @$\wedge$@ (varn env @$\equiv$@ c10 env) @$\rightarrow$@ varn env + vari env @$\equiv$@ c10 env
|
|
215 conv e record { pi1 = refl ; pi2 = refl } = +zero
|
|
216
|
|
217
|
|
218 whileTestPCallwP' : (c : @$\mathbb{N}$@ ) @$\rightarrow$@ Set
|
|
219 whileTestPCallwP' c = whileTestPwP {_} {_} c (@$\lambda$@ env s @$\rightarrow$@ loopPwP' (varn env) env refl (conv env s) ( @$\lambda$@ env s @$\rightarrow$@ vari env @$\equiv$@ c10 env ) ) where
|
|
220 conv : (env : Envc ) @$\rightarrow$@ (vari env @$\equiv$@ 0) @$\wedge$@ (varn env @$\equiv$@ c10 env) @$\rightarrow$@ varn env + vari env @$\equiv$@ c10 env
|
|
221 conv e record { pi1 = refl ; pi2 = refl } = +zero
|
|
222
|
|
223 helperCallwP : (c : @$\mathbb{N}$@) @$\rightarrow$@ whileTestPCallwP' c
|
|
224 helperCallwP c = whileTestPwP {_} {_} c (@$\lambda$@ env s @$\rightarrow$@ loopHelper c (record { c10 = c ; varn = c ; vari = zero }) refl +zero)
|
|
225
|
|
226 --
|
|
227 -- Using imply relation to make soundness explicit
|
|
228 -- termination is shown by induction on varn
|
|
229 --
|
|
230
|
|
231 data _implies_ (A B : Set ) : Set (succ Zero) where
|
|
232 proof : ( A @$\rightarrow$@ B ) @$\rightarrow$@ A implies B
|
|
233
|
|
234 whileTestPSem : (c : @$\mathbb{N}$@) @$\rightarrow$@ whileTestP c ( @$\lambda$@ env @$\rightarrow$@ @$\top$@ implies (whileTestStateP s1 env) )
|
|
235 whileTestPSem c = proof ( @$\lambda$@ _ @$\rightarrow$@ record { pi1 = refl ; pi2 = refl } )
|
|
236
|
|
237 whileTestPSemSound : (c : @$\mathbb{N}$@ ) (output : Envc ) @$\rightarrow$@ output @$\equiv$@ whileTestP c (@$\lambda$@ e @$\rightarrow$@ e) @$\rightarrow$@ @$\top$@ implies ((vari output @$\equiv$@ 0) @$\wedge$@ (varn output @$\equiv$@ c))
|
|
238 whileTestPSemSound c output refl = whileTestPSem c
|
|
239
|
|
240
|
|
241 whileConvPSemSound : {l : Level} @$\rightarrow$@ (input : Envc) @$\rightarrow$@ (whileTestStateP s1 input ) implies (whileTestStateP s2 input)
|
|
242 whileConvPSemSound input = proof @$\lambda$@ x @$\rightarrow$@ (conv input x) where
|
|
243 conv : (env : Envc ) @$\rightarrow$@ (vari env @$\equiv$@ 0) @$\wedge$@ (varn env @$\equiv$@ c10 env) @$\rightarrow$@ varn env + vari env @$\equiv$@ c10 env
|
|
244 conv e record { pi1 = refl ; pi2 = refl } = +zero
|
|
245
|
|
246 loopPP : (n : @$\mathbb{N}$@) @$\rightarrow$@ (input : Envc ) @$\rightarrow$@ (n @$\equiv$@ varn input) @$\rightarrow$@ Envc
|
|
247 loopPP zero input refl = input
|
|
248 loopPP (suc n) input refl =
|
|
249 loopPP n (record input { varn = pred (varn input) ; vari = suc (vari input)}) refl
|
|
250
|
|
251 whileLoopPSem : {l : Level} {t : Set l} @$\rightarrow$@ (input : Envc ) @$\rightarrow$@ whileTestStateP s2 input
|
|
252 @$\rightarrow$@ (next : (output : Envc ) @$\rightarrow$@ (whileTestStateP s2 input ) implies (whileTestStateP s2 output) @$\rightarrow$@ t)
|
|
253 @$\rightarrow$@ (exit : (output : Envc ) @$\rightarrow$@ (whileTestStateP s2 input ) implies (whileTestStateP sf output) @$\rightarrow$@ t) @$\rightarrow$@ t
|
|
254 whileLoopPSem env s next exit with varn env | s
|
|
255 ... | zero | _ = exit env (proof (@$\lambda$@ z @$\rightarrow$@ z))
|
|
256 ... | (suc varn ) | refl = next ( record env { varn = varn ; vari = suc (vari env) } ) (proof @$\lambda$@ x @$\rightarrow$@ +-suc varn (vari env) )
|
|
257
|
|
258 loopPPSem : (input output : Envc ) @$\rightarrow$@ output @$\equiv$@ loopPP (varn input) input refl
|
|
259 @$\rightarrow$@ (whileTestStateP s2 input ) @$\rightarrow$@ (whileTestStateP s2 input ) implies (whileTestStateP sf output)
|
|
260 loopPPSem input output refl s2p = loopPPSemInduct (varn input) input refl refl s2p
|
|
261 where
|
|
262 lem : (n : @$\mathbb{N}$@) @$\rightarrow$@ (env : Envc) @$\rightarrow$@ n + suc (vari env) @$\equiv$@ suc (n + vari env)
|
|
263 lem n env = +-suc (n) (vari env)
|
|
264 loopPPSemInduct : (n : @$\mathbb{N}$@) @$\rightarrow$@ (current : Envc) @$\rightarrow$@ (eq : n @$\equiv$@ varn current) @$\rightarrow$@ (loopeq : output @$\equiv$@ loopPP n current eq)
|
|
265 @$\rightarrow$@ (whileTestStateP s2 current ) @$\rightarrow$@ (whileTestStateP s2 current ) implies (whileTestStateP sf output)
|
|
266 loopPPSemInduct zero current refl loopeq refl rewrite loopeq = proof (@$\lambda$@ x @$\rightarrow$@ refl)
|
|
267 loopPPSemInduct (suc n) current refl loopeq refl rewrite (sym (lem n current)) =
|
|
268 whileLoopPSem current refl
|
|
269 (@$\lambda$@ output x @$\rightarrow$@ loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)
|
|
270 (@$\lambda$@ output x @$\rightarrow$@ loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)
|
|
271
|
|
272 whileLoopPSemSound : {l : Level} @$\rightarrow$@ (input output : Envc )
|
|
273 @$\rightarrow$@ whileTestStateP s2 input
|
|
274 @$\rightarrow$@ output @$\equiv$@ loopPP (varn input) input refl
|
|
275 @$\rightarrow$@ (whileTestStateP s2 input ) implies ( whileTestStateP sf output )
|
|
276 whileLoopPSemSound {l} input output pre eq = loopPPSem input output eq pre
|