Mercurial > hg > Members > ryokka > HoareLogic
view whileTestGears.agda @ 35:7c739972cd26
add loopProof
author | ryokka |
---|---|
date | Fri, 13 Dec 2019 19:23:15 +0900 |
parents | 9caff4e4a402 |
children | 320b765a6424 |
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 )))) data whileTestState (c10 : ℕ ) (env : Env ) : Set where error : whileTestState c10 env state1 : ((vari env) ≡ 0) /\ ((varn env) ≡ c10) → whileTestState c10 env state2 : (varn env + vari env ≡ c10) → whileTestState c10 env finstate : ((vari env) ≡ c10 ) → whileTestState c10 env record Context : Set where field c10 : ℕ whileDG : Env whileCond : whileTestState c10 whileDG 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} ; whileCond = {!!} } ) {-# 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} ; whileCond = {!!} } ) next open import Relation.Nullary open import Relation.Binary {-# TERMINATING #-} whileLoopStep : {l : Level} {t : Set l} → Env → (Code : (e : Env ) → 1 ≤ varn e → t) (Code : (e : Env) → 0 ≡ varn e → t) → t whileLoopStep env next exit with <-cmp 0 (varn env) whileLoopStep env next exit | tri≈ _ eq _ = exit env eq whileLoopStep env next exit | tri< gt _ _ = {!!} where lem : (env : Env) → (1 ≤ varn env) → 1 ≤ (varn env - 1) lem env 1<varn = {!!} -- n が 0 の時 は正しい、 n が1の時正しくない whileLoopStep env next exit | tri> _ _ c = ⊥-elim (m<n⇒n≢0 {varn env} {0} c refl) -- can't happen 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) out init = state1 record { pi1 = refl ; pi2 = refl } {-# TERMINATING #-} whileLoopProof : {l : Level} {t : Set l} → Context → (Code : Context → t) (Code : Context → t) → t whileLoopProof cxt next exit = whileLoopStep (whileDG cxt) ( λ env lt → next record cxt { whileDG = env ; whileCond = {!!} } ) ( λ env eq → exit record cxt { whileDG = env ; whileCond = exitCond env eq } ) where proof5 : (e : Env ) → varn e + vari e ≡ c10 cxt → 0 ≡ varn e → vari e ≡ c10 cxt proof5 record { varn = .0 ; vari = vari } refl refl = refl exitCond : (e : Env ) → 0 ≡ varn e → whileTestState (c10 cxt) e exitCond nenv eq1 with whileCond cxt | inspect whileDG cxt ... | state2 cond | record { eq = eq2 } = finstate ( proof5 nenv {!!} eq1 ) ... | _ | _ = error whileConvProof : {l : Level} {t : Set l} → Context → (Code : Context → t) → t whileConvProof cxt next = next record cxt { whileCond = postCond } where proof4 : (e : Env ) → (vari e ≡ 0) /\ (varn e ≡ c10 cxt) → varn e + vari e ≡ c10 cxt proof4 env p1 = let open ≡-Reasoning in begin varn env + vari env ≡⟨ cong ( λ n → n + vari env ) (pi2 p1 ) ⟩ c10 cxt + vari env ≡⟨ cong ( λ n → c10 cxt + n ) (pi1 p1 ) ⟩ c10 cxt + 0 ≡⟨ +-sym {c10 cxt} {0} ⟩ c10 cxt ∎ postCond : whileTestState (c10 cxt) (whileDG cxt) postCond with whileCond cxt ... | state1 cond = state2 (proof4 (whileDG cxt) cond ) ... | _ = error {-# TERMINATING #-} loop : {l : Level} {t : Set l} → Context → (exit : Context → t) → t loop cxt exit = whileLoopProof cxt (λ cxt → loop cxt exit ) exit loopProof : {l : Level} {t : Set l} {P Pre Post Inv : Context → Set } → (c : Context) → (d : Dec (P c)) → Pre c → (Context → (continue : (c1 : Context) → ¬ (P c1) → Inv c1 → t) → (exit : (c2 : Context) → (P c2) → Post c2 → t) → t) → t loopProof {l} {t} {P} {Pre} {Post} {Inv} cxt (false because proof₁) f = loopProof {!!} {!!} {!!} loopProof {l} {t} {P} {Pre} {Post} {Inv} cxt (true because proof₁) f = {!!} proofWhileGear : (c : ℕ) (cxt : Context) → whileTestProof (record cxt { c10 = c ; whileCond = error }) $ λ cxt → whileConvProof cxt $ λ cxt → loop cxt $ λ cxt → vari (whileDG cxt) ≡ c10 cxt proofWhileGear c cxt = {!!} CodeGear : {l : Level} {t : Set l} → (cont : Set → t) → (exit : Set → t) → t CodeGear = {!!}