view whileTestGears.agda @ 81:0122f980427c

clean up
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Thu, 02 Jan 2020 15:33:49 +0900
parents 148feaa1e346
children 33a6fd61c3e6
line wrap: on
line source

module whileTestGears where

open import Function
open import Data.Nat
open import Data.Bool hiding ( _≟_ ; _≤?_ ; _≤_ ; _<_)
open import Data.Product
open import Level renaming ( suc to succ ; zero to Zero )
open import Relation.Nullary using (¬_; Dec; yes; no)
open import Relation.Binary.PropositionalEquality
open import Agda.Builtin.Unit

open import utilities
open  _/\_

-- original codeGear (with non terminatinng )

record Env : Set (succ Zero) where
  field
    varn : ℕ
    vari : ℕ
open Env 

whileTest : {l : Level} {t : Set l}  → (c10 : ℕ) → (Code : Env → t) → t
whileTest c10 next = next (record {varn = c10 ; vari = 0 } )

{-# TERMINATING #-}
whileLoop : {l : Level} {t : Set l} → Env → (Code : Env → t) → 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 (λ env → whileLoop env (λ env1 → env1 ))

proof1 : whileTest 10 (λ env → whileLoop env (λ e → (vari e) ≡ 10 ))
proof1 = refl

-- codeGear with pre-condtion and post-condition
--
--                                                                              ↓PostCondition
whileTest' : {l : Level} {t : Set l}  →  {c10 :  ℕ } → (Code : (env : Env )  → ((vari env) ≡ 0) /\ ((varn env) ≡ c10) → t) → t
whileTest' {_} {_}  {c10} next = next env proof2
  where
    env : Env 
    env = record {vari = 0 ; varn = c10 }
    proof2 : ((vari env) ≡ 0) /\ ((varn env) ≡ 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} → (env : Env ) → {c10 :  ℕ } → ((varn env) + (vari env) ≡ c10) → (Code : Env  → t) → t
whileLoop' env proof next with  ( suc zero  ≤? (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 ≤ zero → ⊥
      1<0 ()
      proof3 : (suc zero  ≤ (varn  env))  → varn env1 + vari env1 ≡ c10
      proof3 (s≤s lt) with varn  env
      proof3 (s≤s z≤n) | zero = ⊥-elim (1<0 p)
      proof3 (s≤s (z≤n {n'}) ) | suc n =  let open ≡-Reasoning  in
          begin
             n' + (vari env + 1) 
          ≡⟨ cong ( λ z → n' + z ) ( +-sym  {vari env} {1} )  ⟩
             n' + (1 + vari env ) 
          ≡⟨ sym ( +-assoc (n')  1 (vari env) ) ⟩
             (n' + 1) + vari env 
          ≡⟨ cong ( λ z → z + vari env )  +1≡suc  ⟩
             (suc n' ) + vari env 
          ≡⟨⟩
             varn env + vari env
          ≡⟨ proof  ⟩
             c10


-- Condition to Invariant
conversion1 : {l : Level} {t : Set l } → (env : Env ) → {c10 :  ℕ } → ((vari env) ≡ 0) /\ ((varn env) ≡ c10)
               → (Code : (env1 : Env ) → (varn env1 + vari env1 ≡ c10) → t) → t
conversion1 env {c10} p1 next = next env proof4
   where
      proof4 : varn env + vari env ≡ c10
      proof4 = let open ≡-Reasoning  in
          begin
            varn env + vari env
          ≡⟨ cong ( λ n → n + vari env ) (pi2 p1 ) ⟩
            c10 + vari env
          ≡⟨ cong ( λ n → c10 + n ) (pi1 p1 ) ⟩
            c10 + 0
          ≡⟨ +-sym {c10} {0} ⟩
            c10


-- all proofs are connected
proofGears : {c10 :  ℕ } → Set
proofGears {c10} = whileTest' {_} {_} {c10} (λ n p1 →  conversion1 n p1 (λ n1 p2 → whileLoop' n1 p2 (λ n2 →  ( vari n2 ≡ c10 )))) 

-- but we cannot prove the soundness of the last condition
--
-- proofGearsMeta : {c10 :  ℕ } →  proofGears {c10}
-- proofGearsMeta {c10} = {!!} -- net yet done

--
-- codeGear with loop step and closed environment
--

open import Relation.Binary

record Envc : Set (succ Zero) where
  field
    c10 : ℕ
    varn : ℕ
    vari : ℕ
open Envc

whileTestP : {l : Level} {t : Set l} → (c10 : ℕ) → (Code : Envc → t) → t
whileTestP c10 next = next (record {varn = c10 ; vari = 0 ; c10 = c10 } )

whileLoopP : {l : Level} {t : Set l} → Envc → (next : Envc → t) → (exit : Envc → t) → t
whileLoopP env next exit with <-cmp 0 (varn env)
whileLoopP env next exit | tri≈ ¬a b ¬c = exit env
whileLoopP env next exit | tri< a ¬b ¬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} → Envc → (next : Envc → t) → (exit : Envc → t) → t
whileLoopP' env@record { c10 = c10 ; varn = zero ; vari = vari } _ exit = exit env
whileLoopP' record { c10 = c10 ; varn = suc varn1 ; vari = vari } next _ = next (record {c10 = c10 ; varn = varn1 ; vari = suc vari })

-- normal loop without termination
{-# TERMINATING #-}
loopP : {l : Level} {t : Set l} → Envc → (exit : Envc → t) → t
loopP env exit = whileLoopP env (λ env → loopP env exit ) exit

whileTestPCall : (c10 :  ℕ ) → Envc
whileTestPCall c10 = whileTestP {_} {_} c10 (λ env → loopP env (λ env →  env))

--
-- codeGears with states of condition
--
data whileTestState  : Set where
  s1 : whileTestState
  s2 : whileTestState
  sf : whileTestState

whileTestStateP : whileTestState → Envc →  Set 
whileTestStateP s1 env = (vari env ≡ 0) /\ (varn env ≡ c10 env) 
whileTestStateP s2 env = (varn env + vari env ≡ c10 env) 
whileTestStateP sf env = (vari env ≡ c10 env) 

whileTestPwP : {l : Level} {t : Set l} → (c10 : ℕ) → ((env : Envc ) → whileTestStateP s1 env → t) → t
whileTestPwP c10 next = next env record { pi1 = refl ; pi2 = refl } where
   env : Envc
   env = whileTestP c10 ( λ env → env )

whileLoopPwP : {l : Level} {t : Set l}   → (env : Envc ) → whileTestStateP s2 env
    → (next : (env : Envc ) → whileTestStateP s2 env  → t)
    → (exit : (env : Envc ) → whileTestStateP sf env  → t) → t
whileLoopPwP env s next exit with <-cmp 0 (varn env)
whileLoopPwP env s next exit | tri≈ ¬a b ¬c = exit env (lem (sym b) s)
  where
    lem : (varn env ≡ 0) → (varn env + vari env ≡ c10 env) → vari env ≡ c10 env
    lem p1 p2 rewrite p1 = p2
whileLoopPwP env s next exit | tri< a ¬b ¬c  = next (record env {varn = (varn env) - 1 ; vari = (vari env) + 1 }) (proof5 a)
  where
    1<0 : 1 ≤ zero → ⊥
    1<0 ()
    proof5 : (suc zero  ≤ (varn  env))  → (varn env - 1) + (vari env + 1) ≡ c10 env
    proof5 (s≤s lt) with varn  env
    proof5 (s≤s z≤n) | zero = ⊥-elim (1<0 a)
    proof5 (s≤s (z≤n {n'}) ) | suc n = let open ≡-Reasoning in
      begin
        n' + (vari env + 1)
      ≡⟨ cong ( λ z → n' + z ) ( +-sym  {vari env} {1} )  ⟩
        n' + (1 + vari env )
      ≡⟨ sym ( +-assoc (n')  1 (vari env) ) ⟩
        (n' + 1) + vari env
      ≡⟨ cong ( λ z → z + vari env )  +1≡suc  ⟩
        (suc n' ) + vari env
      ≡⟨⟩
        varn env + vari env
      ≡⟨ s  ⟩
         c10 env


{-# TERMINATING #-}
loopPwP : {l : Level} {t : Set l} → (env : Envc ) → whileTestStateP s2 env → (exit : (env : Envc ) → whileTestStateP sf env → t) → t
loopPwP env s exit = whileLoopPwP env s (λ env s → loopPwP env s exit ) exit

--  all codtions are correctly connected and required condtion is proved in the continuation
--      use required condition as t in (env → t) → t 
whileTestPCallwP : (c :  ℕ ) → Set
whileTestPCallwP c = whileTestPwP {_} {_} c ( λ env s → loopPwP env (conv env s) ( λ env s → vari env ≡ c )  ) 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

--
-- 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 = whileTestPSem c

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