view 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
line wrap: on
line source

module whileTestProof where
--
-- 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 → B ) → A implies B

implies2p : {A B : Set } → A implies B  → A → B
implies2p (proof x) = x

whileTestPSem :  (c : ℕ) → whileTestP c ( λ env → ⊤ implies (whileTestStateP s1 env) )
whileTestPSem c = proof ( λ _ → record { pi1 = refl ; pi2 = refl } )

SemGears : (f : {l : Level } {t : Set l } → (e0 : Envc ) → ((e : Envc) → t)  → t ) → Set (succ Zero)
SemGears f = Envc → Envc → Set

GearsUnitSound : (e0 e1 : Envc) {pre : Envc → Set} {post : Envc → Set}
   → (f : {l : Level } {t : Set l } → (e0 : Envc ) → (Envc → t)  → t )
   → (fsem : (e0 : Envc ) → f e0 ( λ e1 → (pre e0) implies (post e1)))
   → f e0 (λ e1 → pre e0 implies post e1)
GearsUnitSound e0 e1 f fsem = fsem e0

whileTestPSemSound : (c : ℕ ) (output : Envc ) → output ≡ whileTestP c (λ e → e) → ⊤ implies ((vari output ≡ 0) /\ (varn output ≡ c))
whileTestPSemSound c output refl = proof (λ x → record { pi1 = refl ; pi2 = refl })
-- whileTestPSem c


whileConvPSemSound : {l : Level} → (input : Envc) → (whileTestStateP s1 input ) implies (whileTestStateP s2 input)
whileConvPSemSound input = proof λ x → (conv input x) where
  conv : (env : Envc ) → (vari env ≡ 0) /\ (varn env ≡ c10 env) → varn env + vari env ≡ c10 env
  conv e record { pi1 = refl ; pi2 = refl } = +zero

loopPP : (n : ℕ) → (input : Envc ) → (n ≡ varn input) → 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}   → (input : Envc ) → whileTestStateP s2 input
  → (next : (output : Envc ) → (whileTestStateP s2 input ) implies (whileTestStateP s2 output)  → t)
  → (exit : (output : Envc ) → (whileTestStateP s2 input ) implies (whileTestStateP sf output)  → t) → t
whileLoopPSem env s next exit with varn env | s
... | zero | _ = exit env (proof (λ z → z))
... | (suc varn ) | refl = next ( record env { varn = varn ; vari = suc (vari env) } ) (proof λ x → +-suc varn (vari env) )

loopPPSem : (input output : Envc ) →  output ≡ loopPP (varn input)  input refl
  → (whileTestStateP s2 input ) → (whileTestStateP s2 input ) implies (whileTestStateP sf output)
loopPPSem input output refl s2p = loopPPSemInduct (varn input) input  refl refl s2p
  where
    lem : (n : ℕ) → (env : Envc) → n + suc (vari env) ≡ suc (n + vari env)
    lem n env = +-suc (n) (vari env)
    loopPPSemInduct : (n : ℕ) → (current : Envc) → (eq : n ≡ varn current) →  (loopeq : output ≡ loopPP n current eq)
      → (whileTestStateP s2 current ) → (whileTestStateP s2 current ) implies (whileTestStateP sf output)
    loopPPSemInduct zero current refl loopeq refl rewrite loopeq = proof (λ x → refl)
    loopPPSemInduct (suc n) current refl loopeq refl rewrite (sym (lem n current)) =
        whileLoopPSem current refl
            (λ output x → loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)
            (λ output x → loopPPSemInduct n (record { c10 = n + suc (vari current) ; varn = n ; vari = suc (vari current) }) refl loopeq refl)

whileLoopPSemSound : {l : Level} → (input output : Envc )
   → whileTestStateP s2 input
   →  output ≡ loopPP (varn input) input refl
   → (whileTestStateP s2 input ) implies ( whileTestStateP sf output )
whileLoopPSemSound {l} input output pre eq = loopPPSem input output eq pre