Mercurial > hg > Papers > 2021 > soto-thesis
diff paper/src/whileTestGears.agda.replaced @ 3:959f4b34d6f4
add final thesis
author | soto |
---|---|
date | Tue, 09 Feb 2021 18:44:53 +0900 |
parents | |
children |
line wrap: on
line diff
--- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/paper/src/whileTestGears.agda.replaced Tue Feb 09 18:44:53 2021 +0900 @@ -0,0 +1,276 @@ +module whileTestGears where + +open import Function +open import Data.Nat +open import Data.Bool hiding ( _@$\stackrel{?}{=}$@_ ; _@$\leq$@?_ ; _@$\leq$@_ ; _<_) +open import Data.Product +open import Level renaming ( suc to succ ; zero to Zero ) +open import Relation.Nullary using (@$\neg$@_; Dec; yes; no) +open import Relation.Binary.PropositionalEquality +open import Agda.Builtin.Unit + +open import utilities +open _@$\wedge$@_ + +-- original codeGear (with non terminatinng ) + +record Env : Set (succ Zero) where + field + varn : @$\mathbb{N}$@ + vari : @$\mathbb{N}$@ +open Env + +whileTest : {l : Level} {t : Set l} @$\rightarrow$@ (c10 : @$\mathbb{N}$@) @$\rightarrow$@ (Code : Env @$\rightarrow$@ t) @$\rightarrow$@ t +whileTest c10 next = next (record {varn = c10 ; vari = 0 } ) + +{-@$\#$@ TERMINATING @$\#$@-} +whileLoop : {l : Level} {t : Set l} @$\rightarrow$@ Env @$\rightarrow$@ (Code : Env @$\rightarrow$@ t) @$\rightarrow$@ t +whileLoop env next with lt 0 (varn env) +whileLoop env next | false = next env +whileLoop env next | true = + whileLoop (record env {varn = (varn env) - 1 ; vari = (vari env) + 1}) next + +test1 : Env +test1 = whileTest 10 (@$\lambda$@ env @$\rightarrow$@ whileLoop env (@$\lambda$@ env1 @$\rightarrow$@ env1 )) + +proof1 : whileTest 10 (@$\lambda$@ env @$\rightarrow$@ whileLoop env (@$\lambda$@ e @$\rightarrow$@ (vari e) @$\equiv$@ 10 )) +proof1 = refl + +-- codeGear with pre-condtion and post-condition +-- +-- ↓PostCondition +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 +whileTest' {_} {_} {c10} next = next env proof2 + where + env : Env + env = record {vari = 0 ; varn = c10 } + proof2 : ((vari env) @$\equiv$@ 0) @$\wedge$@ ((varn env) @$\equiv$@ c10) -- PostCondition + proof2 = record {pi1 = refl ; pi2 = refl} + + +open import Data.Empty +open import Data.Nat.Properties + + +{-@$\#$@ TERMINATING @$\#$@-} -- ↓PreCondition(Invaliant) +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 +whileLoop' env proof next with ( suc zero @$\leq$@? (varn env) ) +whileLoop' env proof next | no p = next env +whileLoop' env {c10} proof next | yes p = whileLoop' env1 (proof3 p ) next + where + env1 = record env {varn = (varn env) - 1 ; vari = (vari env) + 1} + 1<0 : 1 @$\leq$@ zero @$\rightarrow$@ @$\bot$@ + 1<0 () + proof3 : (suc zero @$\leq$@ (varn env)) @$\rightarrow$@ varn env1 + vari env1 @$\equiv$@ c10 + proof3 (s@$\leq$@s lt) with varn env + proof3 (s@$\leq$@s z@$\leq$@n) | zero = @$\bot$@-elim (1<0 p) + proof3 (s@$\leq$@s (z@$\leq$@n {n'}) ) | suc n = let open @$\equiv$@-Reasoning in + begin + n' + (vari env + 1) + @$\equiv$@@$\langle$@ cong ( @$\lambda$@ z @$\rightarrow$@ n' + z ) ( +-sym {vari env} {1} ) @$\rangle$@ + n' + (1 + vari env ) + @$\equiv$@@$\langle$@ sym ( +-assoc (n') 1 (vari env) ) @$\rangle$@ + (n' + 1) + vari env + @$\equiv$@@$\langle$@ cong ( @$\lambda$@ z @$\rightarrow$@ z + vari env ) +1@$\equiv$@suc @$\rangle$@ + (suc n' ) + vari env + @$\equiv$@@$\langle$@@$\rangle$@ + varn env + vari env + @$\equiv$@@$\langle$@ proof @$\rangle$@ + c10 + @$\blacksquare$@ + +-- Condition to Invariant +conversion1 : {l : Level} {t : Set l } @$\rightarrow$@ (env : Env ) @$\rightarrow$@ {c10 : @$\mathbb{N}$@ } @$\rightarrow$@ ((vari env) @$\equiv$@ 0) @$\wedge$@ ((varn env) @$\equiv$@ c10) + @$\rightarrow$@ (Code : (env1 : Env ) @$\rightarrow$@ (varn env1 + vari env1 @$\equiv$@ c10) @$\rightarrow$@ t) @$\rightarrow$@ t +conversion1 env {c10} p1 next = next env proof4 + where + proof4 : varn env + vari env @$\equiv$@ c10 + proof4 = let open @$\equiv$@-Reasoning in + begin + varn env + vari env + @$\equiv$@@$\langle$@ cong ( @$\lambda$@ n @$\rightarrow$@ n + vari env ) (pi2 p1 ) @$\rangle$@ + c10 + vari env + @$\equiv$@@$\langle$@ cong ( @$\lambda$@ n @$\rightarrow$@ c10 + n ) (pi1 p1 ) @$\rangle$@ + c10 + 0 + @$\equiv$@@$\langle$@ +-sym {c10} {0} @$\rangle$@ + c10 + @$\blacksquare$@ + +-- all proofs are connected +proofGears : {c10 : @$\mathbb{N}$@ } @$\rightarrow$@ Set +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 )))) + +-- +-- codeGear with loop step and closed environment +-- + +open import Relation.Binary + +record Envc : Set (succ Zero) where + field + c10 : @$\mathbb{N}$@ + varn : @$\mathbb{N}$@ + vari : @$\mathbb{N}$@ +open Envc + +whileTestP : {l : Level} {t : Set l} @$\rightarrow$@ (c10 : @$\mathbb{N}$@) @$\rightarrow$@ (Code : Envc @$\rightarrow$@ t) @$\rightarrow$@ t +whileTestP c10 next = next (record {varn = c10 ; vari = 0 ; c10 = c10 } ) + +whileLoopP : {l : Level} {t : Set l} @$\rightarrow$@ Envc @$\rightarrow$@ (next : Envc @$\rightarrow$@ t) @$\rightarrow$@ (exit : Envc @$\rightarrow$@ t) @$\rightarrow$@ t +whileLoopP env next exit with <-cmp 0 (varn env) +whileLoopP env next exit | tri@$\thickapprox$@ @$\neg$@a b @$\neg$@c = exit env +whileLoopP env next exit | tri< a @$\neg$@b @$\neg$@c = + next (record env {varn = (varn env) - 1 ; vari = (vari env) + 1 }) + +-- equivalent of whileLoopP but it looks like an induction on varn +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 +whileLoopP' zero env refl _ exit = exit env +whileLoopP' (suc n) env refl next _ = next (record {c10 = (c10 env) ; varn = varn env ; vari = suc (vari env) }) + +-- normal loop without termination +{-@$\#$@ TERMINATING @$\#$@-} +loopP : {l : Level} {t : Set l} @$\rightarrow$@ Envc @$\rightarrow$@ (exit : Envc @$\rightarrow$@ t) @$\rightarrow$@ t +loopP env exit = whileLoopP env (@$\lambda$@ env @$\rightarrow$@ loopP env exit ) exit + +whileTestPCall : (c10 : @$\mathbb{N}$@ ) @$\rightarrow$@ Envc +whileTestPCall c10 = whileTestP {_} {_} c10 (@$\lambda$@ env @$\rightarrow$@ loopP env (@$\lambda$@ env @$\rightarrow$@ env)) + +-- +-- codeGears with states of condition +-- +data whileTestState : Set where + s1 : whileTestState + s2 : whileTestState + sf : whileTestState + +whileTestStateP : whileTestState @$\rightarrow$@ Envc @$\rightarrow$@ Set +whileTestStateP s1 env = (vari env @$\equiv$@ 0) @$\wedge$@ (varn env @$\equiv$@ c10 env) +whileTestStateP s2 env = (varn env + vari env @$\equiv$@ c10 env) +whileTestStateP sf env = (vari env @$\equiv$@ c10 env) + +whileTestPwP : {l : Level} {t : Set l} @$\rightarrow$@ (c10 : @$\mathbb{N}$@) @$\rightarrow$@ ((env : Envc ) @$\rightarrow$@ whileTestStateP s1 env @$\rightarrow$@ t) @$\rightarrow$@ t +whileTestPwP c10 next = next env record { pi1 = refl ; pi2 = refl } where + env : Envc + env = whileTestP c10 ( @$\lambda$@ env @$\rightarrow$@ env ) + +whileLoopPwP : {l : Level} {t : Set l} @$\rightarrow$@ (env : Envc ) @$\rightarrow$@ whileTestStateP s2 env + @$\rightarrow$@ (next : (env : Envc ) @$\rightarrow$@ whileTestStateP s2 env @$\rightarrow$@ t) + @$\rightarrow$@ (exit : (env : Envc ) @$\rightarrow$@ whileTestStateP sf env @$\rightarrow$@ t) @$\rightarrow$@ t +whileLoopPwP env s next exit with <-cmp 0 (varn env) +whileLoopPwP env s next exit | tri@$\thickapprox$@ @$\neg$@a b @$\neg$@c = exit env (lem (sym b) s) + where + lem : (varn env @$\equiv$@ 0) @$\rightarrow$@ (varn env + vari env @$\equiv$@ c10 env) @$\rightarrow$@ vari env @$\equiv$@ c10 env + lem refl refl = refl +whileLoopPwP env s next exit | tri< a @$\neg$@b @$\neg$@c = next (record env {varn = (varn env) - 1 ; vari = (vari env) + 1 }) (proof5 a) + where + 1<0 : 1 @$\leq$@ zero @$\rightarrow$@ @$\bot$@ + 1<0 () + proof5 : (suc zero @$\leq$@ (varn env)) @$\rightarrow$@ (varn env - 1) + (vari env + 1) @$\equiv$@ c10 env + proof5 (s@$\leq$@s lt) with varn env + proof5 (s@$\leq$@s z@$\leq$@n) | zero = @$\bot$@-elim (1<0 a) + proof5 (s@$\leq$@s (z@$\leq$@n {n'}) ) | suc n = let open @$\equiv$@-Reasoning in + begin + n' + (vari env + 1) + @$\equiv$@@$\langle$@ cong ( @$\lambda$@ z @$\rightarrow$@ n' + z ) ( +-sym {vari env} {1} ) @$\rangle$@ + n' + (1 + vari env ) + @$\equiv$@@$\langle$@ sym ( +-assoc (n') 1 (vari env) ) @$\rangle$@ + (n' + 1) + vari env + @$\equiv$@@$\langle$@ cong ( @$\lambda$@ z @$\rightarrow$@ z + vari env ) +1@$\equiv$@suc @$\rangle$@ + (suc n' ) + vari env + @$\equiv$@@$\langle$@@$\rangle$@ + varn env + vari env + @$\equiv$@@$\langle$@ s @$\rangle$@ + c10 env + @$\blacksquare$@ + + +whileLoopPwP' : {l : Level} {t : Set l} @$\rightarrow$@ (n : @$\mathbb{N}$@) @$\rightarrow$@ (env : Envc ) @$\rightarrow$@ (n @$\equiv$@ varn env) @$\rightarrow$@ whileTestStateP s2 env + @$\rightarrow$@ (next : (env : Envc ) @$\rightarrow$@ (pred n @$\equiv$@ varn env) @$\rightarrow$@ whileTestStateP s2 env @$\rightarrow$@ t) + @$\rightarrow$@ (exit : (env : Envc ) @$\rightarrow$@ whileTestStateP sf env @$\rightarrow$@ t) @$\rightarrow$@ t +whileLoopPwP' zero env refl refl next exit = exit env refl +whileLoopPwP' (suc n) env refl refl next exit = next (record env {varn = pred (varn env) ; vari = suc (vari env) }) refl (+-suc n (vari env)) + + +{-@$\#$@ TERMINATING @$\#$@-} +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 +loopPwP env s exit = whileLoopPwP env s (@$\lambda$@ env s @$\rightarrow$@ loopPwP env s exit ) exit + + +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 +loopPwP' zero env refl refl exit = exit env refl +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 + + +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}$@) +loopHelper zero env eq refl rewrite eq = refl +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)) + + +-- all codtions are correctly connected and required condtion is proved in the continuation +-- use required condition as t in (env @$\rightarrow$@ t) @$\rightarrow$@ t +-- +whileTestPCallwP : (c : @$\mathbb{N}$@ ) @$\rightarrow$@ Set +whileTestPCallwP c = whileTestPwP {_} {_} c ( @$\lambda$@ env s @$\rightarrow$@ loopPwP env (conv env s) ( @$\lambda$@ env s @$\rightarrow$@ vari env @$\equiv$@ c10 env ) ) where + conv : (env : Envc ) @$\rightarrow$@ (vari env @$\equiv$@ 0) @$\wedge$@ (varn env @$\equiv$@ c10 env) @$\rightarrow$@ varn env + vari env @$\equiv$@ c10 env + conv e record { pi1 = refl ; pi2 = refl } = +zero + + +whileTestPCallwP' : (c : @$\mathbb{N}$@ ) @$\rightarrow$@ Set +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 + conv : (env : Envc ) @$\rightarrow$@ (vari env @$\equiv$@ 0) @$\wedge$@ (varn env @$\equiv$@ c10 env) @$\rightarrow$@ varn env + vari env @$\equiv$@ c10 env + conv e record { pi1 = refl ; pi2 = refl } = +zero + +helperCallwP : (c : @$\mathbb{N}$@) @$\rightarrow$@ whileTestPCallwP' c +helperCallwP c = whileTestPwP {_} {_} c (@$\lambda$@ env s @$\rightarrow$@ loopHelper c (record { c10 = c ; varn = c ; vari = zero }) refl +zero) + +-- +-- Using imply relation to make soundness explicit +-- termination is shown by induction on varn +-- + +data _implies_ (A B : Set ) : Set (succ Zero) where + proof : ( A @$\rightarrow$@ B ) @$\rightarrow$@ A implies B + +whileTestPSem : (c : @$\mathbb{N}$@) @$\rightarrow$@ whileTestP c ( @$\lambda$@ env @$\rightarrow$@ @$\top$@ implies (whileTestStateP s1 env) ) +whileTestPSem c = proof ( @$\lambda$@ _ @$\rightarrow$@ record { pi1 = refl ; pi2 = refl } ) + +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)) +whileTestPSemSound c output refl = whileTestPSem c + + +whileConvPSemSound : {l : Level} @$\rightarrow$@ (input : Envc) @$\rightarrow$@ (whileTestStateP s1 input ) implies (whileTestStateP s2 input) +whileConvPSemSound input = proof @$\lambda$@ x @$\rightarrow$@ (conv input x) where + conv : (env : Envc ) @$\rightarrow$@ (vari env @$\equiv$@ 0) @$\wedge$@ (varn env @$\equiv$@ c10 env) @$\rightarrow$@ varn env + vari env @$\equiv$@ c10 env + conv e record { pi1 = refl ; pi2 = refl } = +zero + +loopPP : (n : @$\mathbb{N}$@) @$\rightarrow$@ (input : Envc ) @$\rightarrow$@ (n @$\equiv$@ varn input) @$\rightarrow$@ Envc +loopPP zero input refl = input +loopPP (suc n) input refl = + loopPP n (record input { varn = pred (varn input) ; vari = suc (vari input)}) refl + +whileLoopPSem : {l : Level} {t : Set l} @$\rightarrow$@ (input : Envc ) @$\rightarrow$@ whileTestStateP s2 input + @$\rightarrow$@ (next : (output : Envc ) @$\rightarrow$@ (whileTestStateP s2 input ) implies (whileTestStateP s2 output) @$\rightarrow$@ t) + @$\rightarrow$@ (exit : (output : Envc ) @$\rightarrow$@ (whileTestStateP s2 input ) implies (whileTestStateP sf output) @$\rightarrow$@ t) @$\rightarrow$@ t +whileLoopPSem env s next exit with varn env | s +... | zero | _ = exit env (proof (@$\lambda$@ z @$\rightarrow$@ z)) +... | (suc varn ) | refl = next ( record env { varn = varn ; vari = suc (vari env) } ) (proof @$\lambda$@ x @$\rightarrow$@ +-suc varn (vari env) ) + +loopPPSem : (input output : Envc ) @$\rightarrow$@ output @$\equiv$@ loopPP (varn input) input refl + @$\rightarrow$@ (whileTestStateP s2 input ) @$\rightarrow$@ (whileTestStateP s2 input ) implies (whileTestStateP sf output) +loopPPSem input output refl s2p = loopPPSemInduct (varn input) input refl refl s2p + where + lem : (n : @$\mathbb{N}$@) @$\rightarrow$@ (env : Envc) @$\rightarrow$@ n + suc (vari env) @$\equiv$@ suc (n + vari env) + lem n env = +-suc (n) (vari env) + loopPPSemInduct : (n : @$\mathbb{N}$@) @$\rightarrow$@ (current : Envc) @$\rightarrow$@ (eq : n @$\equiv$@ varn current) @$\rightarrow$@ (loopeq : output @$\equiv$@ loopPP n current eq) + @$\rightarrow$@ (whileTestStateP s2 current ) @$\rightarrow$@ (whileTestStateP s2 current ) implies (whileTestStateP sf output) + loopPPSemInduct zero current refl loopeq refl rewrite loopeq = proof (@$\lambda$@ x @$\rightarrow$@ refl) + loopPPSemInduct (suc n) current refl loopeq refl rewrite (sym (lem n current)) = + whileLoopPSem current refl + (@$\lambda$@ output x @$\rightarrow$@ loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl) + (@$\lambda$@ output x @$\rightarrow$@ loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl) + +whileLoopPSemSound : {l : Level} @$\rightarrow$@ (input output : Envc ) + @$\rightarrow$@ whileTestStateP s2 input + @$\rightarrow$@ output @$\equiv$@ loopPP (varn input) input refl + @$\rightarrow$@ (whileTestStateP s2 input ) implies ( whileTestStateP sf output ) +whileLoopPSemSound {l} input output pre eq = loopPPSem input output eq pre