comparison src/whileTestProof.agda @ 1:73127e0ab57c

(none)
author soto@cr.ie.u-ryukyu.ac.jp
date Tue, 08 Sep 2020 18:38:08 +0900
parents
children
comparison
equal deleted inserted replaced
0:b919985837a3 1:73127e0ab57c
1 module whileTestProof where
2 --
3 -- Using imply relation to make soundness explicit
4 -- termination is shown by induction on varn
5 --
6
7 data _implies_ (A B : Set ) : Set (succ Zero) where
8 proof : ( A → B ) → A implies B
9
10 implies2p : {A B : Set } → A implies B → A → B
11 implies2p (proof x) = x
12
13 whileTestPSem : (c : ℕ) → whileTestP c ( λ env → ⊤ implies (whileTestStateP s1 env) )
14 whileTestPSem c = proof ( λ _ → record { pi1 = refl ; pi2 = refl } )
15
16 SemGears : (f : {l : Level } {t : Set l } → (e0 : Envc ) → ((e : Envc) → t) → t ) → Set (succ Zero)
17 SemGears f = Envc → Envc → Set
18
19 GearsUnitSound : (e0 e1 : Envc) {pre : Envc → Set} {post : Envc → Set}
20 → (f : {l : Level } {t : Set l } → (e0 : Envc ) → (Envc → t) → t )
21 → (fsem : (e0 : Envc ) → f e0 ( λ e1 → (pre e0) implies (post e1)))
22 → f e0 (λ e1 → pre e0 implies post e1)
23 GearsUnitSound e0 e1 f fsem = fsem e0
24
25 whileTestPSemSound : (c : ℕ ) (output : Envc ) → output ≡ whileTestP c (λ e → e) → ⊤ implies ((vari output ≡ 0) /\ (varn output ≡ c))
26 whileTestPSemSound c output refl = proof (λ x → record { pi1 = refl ; pi2 = refl })
27 -- whileTestPSem c
28
29
30 whileConvPSemSound : {l : Level} → (input : Envc) → (whileTestStateP s1 input ) implies (whileTestStateP s2 input)
31 whileConvPSemSound input = proof λ x → (conv input x) where
32 conv : (env : Envc ) → (vari env ≡ 0) /\ (varn env ≡ c10 env) → varn env + vari env ≡ c10 env
33 conv e record { pi1 = refl ; pi2 = refl } = +zero
34
35 loopPP : (n : ℕ) → (input : Envc ) → (n ≡ varn input) → Envc
36 loopPP zero input refl = input
37 loopPP (suc n) input refl =
38 loopPP n (record input { varn = pred (varn input) ; vari = suc (vari input)}) refl
39
40 whileLoopPSem : {l : Level} {t : Set l} → (input : Envc ) → whileTestStateP s2 input
41 → (next : (output : Envc ) → (whileTestStateP s2 input ) implies (whileTestStateP s2 output) → t)
42 → (exit : (output : Envc ) → (whileTestStateP s2 input ) implies (whileTestStateP sf output) → t) → t
43 whileLoopPSem env s next exit with varn env | s
44 ... | zero | _ = exit env (proof (λ z → z))
45 ... | (suc varn ) | refl = next ( record env { varn = varn ; vari = suc (vari env) } ) (proof λ x → +-suc varn (vari env) )
46
47 loopPPSem : (input output : Envc ) → output ≡ loopPP (varn input) input refl
48 → (whileTestStateP s2 input ) → (whileTestStateP s2 input ) implies (whileTestStateP sf output)
49 loopPPSem input output refl s2p = loopPPSemInduct (varn input) input refl refl s2p
50 where
51 lem : (n : ℕ) → (env : Envc) → n + suc (vari env) ≡ suc (n + vari env)
52 lem n env = +-suc (n) (vari env)
53 loopPPSemInduct : (n : ℕ) → (current : Envc) → (eq : n ≡ varn current) → (loopeq : output ≡ loopPP n current eq)
54 → (whileTestStateP s2 current ) → (whileTestStateP s2 current ) implies (whileTestStateP sf output)
55 loopPPSemInduct zero current refl loopeq refl rewrite loopeq = proof (λ x → refl)
56 loopPPSemInduct (suc n) current refl loopeq refl rewrite (sym (lem n current)) =
57 whileLoopPSem current refl
58 (λ output x → loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)
59 (λ output x → loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)
60
61 whileLoopPSemSound : {l : Level} → (input output : Envc )
62 → whileTestStateP s2 input
63 → output ≡ loopPP (varn input) input refl
64 → (whileTestStateP s2 input ) implies ( whileTestStateP sf output )
65 whileLoopPSemSound {l} input output pre eq = loopPPSem input output eq pre