view whileTestGears.agda @ 30:dd66b94bf365

loop causes agda inifinite loop
author Shinji KONO <kono@ie.u-ryukyu.ac.jp>
date Tue, 10 Dec 2019 08:57:11 +0900
parents 816b44e5e674
children 600b4e914071
line wrap: on
line source

module whileTestGears where

open import Function
open import Data.Nat
open import Data.Bool hiding ( _≟_ ; _≤?_ ; _≤_)
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 utilities
open  _/\_

record Env  : Set 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 {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 {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 Invaliant
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 :  ℕ } → whileTest' {_} {_} {c10} (λ n p1 →  conversion1 n p1 (λ n1 p2 → whileLoop' n1 p2 (λ n2 →  ( vari n2 ≡ c10 )))) 
-- proofGearsMeta {c10} = {!!}

data whileTestState (c10 : ℕ ) : Set where
  error : whileTestState c10
  state1 : (env : Env)  → ((vari env) ≡ 0) /\ ((varn env) ≡ c10) → whileTestState c10
  state2 : (env : Env) → (varn env + vari env ≡ c10) → whileTestState c10
  finstate : (env : Env)  → ((vari env) ≡ c10 ) → whileTestState c10

whileLoopProof0 : {c10 : ℕ } → Set
whileLoopProof0 {c10 } = Env /\ whileTestState c10 → whileLoop {!!} ( λ env → vari env ≡ c10 )  
whileTestProof0 : {c10 : ℕ } → Set
whileTestProof0 {c10} = Env /\ whileTestState c10 → whileTest c10 (λ env →  {!!} )

proofGearsState : {c10 :  ℕ } → whileTestProof0
proofGearsState {c10} = {!!}  where 
  lemma1 : (env : Env ) →  vari env ≡ c10
  lemma1 = {!!}

record Context : Set where
   field 
     c10 : ℕ
     whileDG : Env
     whileCond : whileTestState c10

open Context

whileTestContext : {l : Level} {t : Set l} -> Context → (Code : Context -> t) -> t
whileTestContext cxt next = next (record cxt { whileDG = record (whileDG cxt) {varn = c10 cxt ; vari = 0}} )

{-# TERMINATING #-}
whileLoopContext : {l : Level} {t : Set l} -> Context -> (Code : Context -> t) -> t
whileLoopContext cxt next with lt 0 (varn (whileDG cxt) )
whileLoopContext cxt next | false = next cxt
whileLoopContext cxt next | true =
    whileLoopContext (record cxt { whileDG = record {varn = (varn (whileDG cxt)) - 1 ; vari = (vari (whileDG cxt)) + 1} } )  next

{-# TERMINATING #-}
whileLoopStep : {l : Level} {t : Set l} -> Env -> (Code : Env -> t) (Code : Env -> t) -> t
whileLoopStep env next exit with lt 0 (varn env)
whileLoopStep env next exit | false = exit env
whileLoopStep env next exit | true =
    next (record {varn = (varn env) - 1 ; vari = (vari env) + 1}) 

whileTestProof : {l : Level} {t : Set l} -> Context → (Code : Context -> t) -> t
whileTestProof cxt next = next record cxt { whileDG = out ; whileCond = init } where
    out : Env 
    out =  whileTest (c10 cxt) ( λ e → e ) 
    init : whileTestState (c10 cxt)
    init = state1 out record { pi1 = refl ; pi2 = refl }

{-# TERMINATING #-}
whileLoopProof : {l : Level} {t : Set l} -> Context -> (Code : Context -> t) (Code : Context -> t) -> t
whileLoopProof cxt next exit = next record cxt { whileDG = out ; whileCond = postCond }  where
    out : Env 
    out =  whileLoopStep (whileDG cxt) ( λ e → e ) {!!}
    postCond : whileTestState (c10 cxt) 
    postCond with whileCond cxt
    ... | state2 e inv = state2 out {!!} where
       proof3 : {!!} 
       proof3 = {!!}
    ... | _ = error

whileConvProof : {l : Level} {t : Set l} -> Context -> (Code : Context -> t) -> t
whileConvProof cxt next = next record cxt { whileCond = postCond } where
    postCond : whileTestState (c10 cxt)
    postCond with whileCond cxt
    ... | state1 e inv = state2 (whileDG cxt) {!!} 
    ... | _ = error

whileFinConvProof : {l : Level} {t : Set l} -> Context -> (Code : Context -> t) -> t
whileFinConvProof cxt next = next record cxt { whileCond = postCond } where
    postCond : whileTestState (c10 cxt)
    postCond with whileCond cxt
    ... | state2 e inv = finstate (whileDG cxt) {!!} 
    ... | _ = error

{-# TERMINATING #-}
loop : {l : Level} {t : Set l} -> Context -> (exit : Context -> t) -> t
loop cxt exit = whileLoopProof cxt (λ cxt → loop cxt exit) exit

proofWhileGear : (c : ℕ) (cxt : Context) → whileTestProof (record cxt { c10 = c ; whileCond = error })
    $ λ cxt → whileConvProof cxt 
    $ λ cxt → loop cxt 
    $ λ cxt → whileFinConvProof cxt
    $ λ cxt → vari (whileDG cxt) ≡ c10 cxt
proofWhileGear cxt = refl