view whileTestGears.agda @ 74:6f26de2fb7fe

...
author ryokka
date Thu, 26 Dec 2019 18:16:35 +0900
parents 52acd110df18
children d45e98211d7d
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  _/\_

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

--                                                                              ↓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



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

-- proofGearsMeta : {c10 :  ℕ } →  proofGears {c10}
-- proofGearsMeta {c10} = {!!} -- net yet done

--
--      openended Env c  <=>  Context
--

open import Relation.Nullary hiding (proof)
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 })

whileLoopP' : {l : Level} {t : Set l} → Envc → (next : Envc → t) → (exit : Envc → t) → t
whileLoopP' env@record { c10 = c10 ; varn = zero ; vari = vari } next exit = exit env
whileLoopP' env@record { c10 = c10 ; varn = (suc varn1) ; vari = vari } next exit = next (record env {varn = varn1 ; vari = vari + 1 })

-- 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 })

{-# 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))

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


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

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 <-cmp 0 (varn env)
whileLoopPSem env s next exit | tri≈ ¬a b ¬c rewrite (sym b) = exit env (proof (λ z → z))
whileLoopPSem env s next exit | tri< a ¬b ¬c  = next env (proof (λ z → z))



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@(record { c10 = c10 ; varn = zero ; vari = vari }) s next exit = exit env (proof (λ z → z))
whileLoopPSem' env@(record { c10 = c10 ; varn = (suc varn₁) ; vari = vari }) s next exit = next env (proof (λ z → z))


loopPP : (n : ℕ) → (input : Envc ) → (n ≡ varn input) → Envc
loopPP n input@(record { c10 = c10 ; varn = zero ; vari = vari }) refl = input
loopPP n input@(record { c10 = c10 ; varn = (suc varn₁) ; vari = vari }) refl = whileLoopP (record { c10 = c10 ; varn = (varn₁) ; vari = vari }) (λ x → loopPP (n - 1)  (record { c10 = c10 ; varn = (varn₁) ; vari = vari }) refl) (λ output → output)

loopPPSem : {l : Level} {t : Set l} → (input output : Envc ) →  output ≡ loopPP (varn input)  input refl
  → (whileTestStateP s2 input ) → (whileTestStateP s2 input ) implies (whileTestStateP sf output)
loopPPSem {l} {t} input output refl s2p = loopPPSemInduct (varn input) input refl s2p (proof (λ x → {!!}))
  where
    -- lem1 :  (ncurrent : Envc) → suc (varn ncurrent) + vari ncurrent ≡ c10 ncurrent → varn ncurrent + suc (vari ncurrent) ≡ c10 ncurrent
    -- lem1 current eq = {!!}
    loopPPSemInduct : (n : ℕ) → (current : Envc) → varn current ≡ n → (whileTestStateP s2 current ) → (whileTestStateP s2 input ) implies (whileTestStateP s2 output) → (whileTestStateP s2 current ) implies (whileTestStateP sf output)
    loopPPSemInduct zero current refl refl imp with loopPP (varn current) current refl -- = proof λ x → {!!}
    loopPPSemInduct zero record { c10 = _ ; varn = _ ; vari = _ } refl refl imp | record { c10 = c10 ; varn = varn ; vari = vari } = proof {!!}
    loopPPSemInduct (suc n) current refl s2p imp with <-cmp 0 (suc n) | whileLoopPSem current s2p (λ output x → proof {!!}) λ output x → proof λ x₁ → {!!}
    loopPPSemInduct (suc n) record { c10 = _ ; varn = .(suc n) ; vari = _ } refl s2p imp | tri< a ¬b ¬c | proof x = {!!}



lpc : (input : Envc ) → Envc
lpc input@(record { c10 = c10 ; varn = zero ; vari = vari }) = input
lpc input@(record { c10 = c10 ; varn = (suc varn₁) ; vari = vari }) = whileLoopP (record { c10 = c10 ; varn = (varn₁) ; vari = vari }) (λ x → lpc record { c10 = c10 ; varn = (varn₁) ; vari = vari }) (λ output → output)




-- loopPP' input | tri≈ ¬a b ¬c = input
-- loopPP' input | tri< a ¬b ¬c = whileLoopP input (λ enext → loopPP' enext) (λ eout → eout)
-- loopPP' (whileLoopP input (λ next → loopPP' next) (λ output → output))


-- = whileLoopP input (λ next → loopPP next ) (λ output → output )

whileLoopPSemSound : (input output : Envc )
   → whileTestStateP s2 input
   →  output ≡ lpc input
   → (whileTestStateP s2 input ) implies ( whileTestStateP sf output )
whileLoopPSemSound input output pre eq = {!!}

-- with (lpc input)
-- record { c10 = c11 ; varn = varn₁ ; vari = vari₁ } .(lpc (record { c10 = c11 ; varn = varn₁ ; vari = vari₁ })) pre refl | record { c10 = c10 ; varn = varn ; vari = vari } = proof λ x → {!!}
  -- where
  --   lem : (whileTestStateP s2 input ) → (varn input + vari input ≡ c10 input)
  --     implies (vari output ≡ c10 output)
  --   lem refl = proof λ x → {!!}


-- whileLoopPSem' input pre (λ output1 x → proof (λ x₁ → ?)) (λ output₁ x → proof (λ x₁ → ?))))
-- proof (whileLoopPwP input pre (λ e p1 p11 → {!!}) (λ e2 p2 p22 → {!!}) )
-- with <-cmp 0 (varn input )
-- ... | ttt = {!!}

-- induction にする
{-# 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

--  wP を Env のRel にする  Env → Env → Set にしちゃう
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


conv1 : (env : Envc ) → (vari env ≡ 0) /\ (varn env ≡ c10 env) → varn env + vari env ≡ c10 env
conv1 e record { pi1 = refl ; pi2 = refl } = +zero

-- = whileTestPwP (suc c) (λ env s → loopPwP env (conv1 env s) (λ env₁ s₁ → {!!}))


data GComm : Set (succ Zero) where
  Skip  : GComm
  Abort : GComm
  PComm : Set → GComm
  -- Seq   : GComm → GComm → GComm
  -- If   : whileTestState → GComm → GComm → GComm
  while : whileTestState → GComm → GComm

gearsSem : {l : Level} {t : Set l} → {c10 : ℕ} → Envc → Envc → (Envc → (Envc → t) → t) → Set
gearsSem pre post = {!!}

unionInf : ∀ {l} -> (ℕ -> Rel Set l) -> Rel Set l
unionInf f a b = ∃ (λ (n : ℕ) → f n a b)

comp : ∀ {l} → Rel Set l → Rel Set l → Rel Set (succ Zero Level.⊔ l)
comp r1 r2 a b = ∃ (λ (a' : Set) → r1 a a' × r2 a' b)

-- repeat : ℕ -> rel set zero -> rel set zero
-- repeat ℕ.zero r = λ x x₁ → ⊤
-- repeat (ℕ.suc m) r = comp (repeat m r) r

GSemComm : {l : Level} {t : Set l} → GComm → Rel whileTestState (Zero)
GSemComm Skip = λ x x₁ → ⊤
GSemComm Abort = λ x x₁ → ⊥
GSemComm (PComm x) = λ x₁ x₂ → x
-- GSemComm (Seq con con₁ con₃) = λ x₁ x₂ → {!!}
-- GSemComm (If x con con₁) = {!!}
GSemComm (while x con) = λ x₁ x₂ → unionInf {Zero} (λ (n : ℕ) →  {!!}) {!!} {!!}

ProofConnect : {l : Level} {t : Set l}
  → (pr1 : Envc → Set → Set)
  → (Envc → Set → (Envc → Set → t))
  → (Envc → Set → Set)
ProofConnect prev f env post =  {!!} -- with f env ({!!}) {!!} 

Proof2 : (env : Envc) → (vari env ≡ c10 env) → vari env ≡ c10 env
Proof2 _ refl = refl 


-- Proof1 : (env : Envc) → (s : varn env + vari env ≡ c10 env) → ((env : Envc) → (vari env ≡ c10 env) → vari env ≡ c10 env) → vari env ≡ c10 env
Proof1 : (env : Envc) → (s : varn env + vari env ≡ c10 env) → loopPwP env s ( λ env s → vari env ≡ c10 env ) 
Proof1 env s = {!!} 

Proof : (c :  ℕ ) → whileTestPCallwP c
Proof c = {!!}