Mercurial > hg > Members > Moririn
view src/parallel_execution/stack.agda @ 179:b3be97ba0782
Prove n-push/n-pop
author | atton |
---|---|
date | Tue, 06 Dec 2016 08:17:08 +0000 |
parents | bf26f1105862 |
children | d8947747ff3b |
line wrap: on
line source
module stack where open import Relation.Binary.PropositionalEquality data Nat : Set where zero : Nat suc : Nat -> Nat data Bool : Set where True : Bool False : Bool data Maybe (a : Set) : Set where Nothing : Maybe a Just : a -> Maybe a record Stack {a t : Set} (stackImpl : Set) : Set where field stack : stackImpl push : stackImpl -> a -> (stackImpl -> t) -> t pop : stackImpl -> (stackImpl -> Maybe a -> t) -> t data Element (a : Set) : Set where cons : a -> Maybe (Element a) -> Element a datum : {a : Set} -> Element a -> a datum (cons a _) = a next : {a : Set} -> Element a -> Maybe (Element a) next (cons _ n) = n {- -- cannot define recrusive record definition. so use linked list with maybe. record Element {l : Level} (a : Set l) : Set (suc l) where field datum : a -- `data` is reserved by Agda. next : Maybe (Element a) -} record SingleLinkedStack (a : Set) : Set where field top : Maybe (Element a) open SingleLinkedStack pushSingleLinkedStack : {Data t : Set} -> SingleLinkedStack Data -> Data -> (Code : SingleLinkedStack Data -> t) -> t pushSingleLinkedStack stack datum next = next stack1 where element = cons datum (top stack) stack1 = record {top = Just element} popSingleLinkedStack : {a t : Set} -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> t) -> t popSingleLinkedStack stack cs with (top stack) ... | Nothing = cs stack Nothing ... | Just d = cs stack1 (Just data1) where data1 = datum d stack1 = record { top = (next d) } emptySingleLinkedStack : {a : Set} -> SingleLinkedStack a emptySingleLinkedStack = record {top = Nothing} createSingleLinkedStack : {a b : Set} -> Stack {a} {b} (SingleLinkedStack a) createSingleLinkedStack = record { stack = emptySingleLinkedStack ; push = pushSingleLinkedStack ; pop = popSingleLinkedStack } test01 : {a : Set} -> SingleLinkedStack a -> Maybe a -> Bool test01 stack _ with (top stack) ... | (Just _) = True ... | Nothing = False test02 : {a : Set} -> SingleLinkedStack a -> Bool test02 stack = (popSingleLinkedStack stack) test01 test03 : {a : Set} -> a -> Bool test03 v = pushSingleLinkedStack emptySingleLinkedStack v test02 lemma : {A : Set} {a : A} -> test03 a ≡ False lemma = refl id : {A : Set} -> A -> A id a = a n-push : {A : Set} {a : A} -> Nat -> SingleLinkedStack A -> SingleLinkedStack A n-push zero s = s n-push {A} {a} (suc n) s = pushSingleLinkedStack (n-push {A} {a} n s) a (\s -> s) n-pop : {A : Set} {a : A} -> Nat -> SingleLinkedStack A -> SingleLinkedStack A n-pop zero s = s n-pop {A} {a} (suc n) s = popSingleLinkedStack (n-pop {A} {a} n s) (\s _ -> s) open ≡-Reasoning push-pop-equiv : {A : Set} {a : A} (s : SingleLinkedStack A) -> popSingleLinkedStack (pushSingleLinkedStack s a (\s -> s)) (\s _ -> s) ≡ s push-pop-equiv s = refl push-and-n-pop : {A : Set} {a : A} (n : Nat) (s : SingleLinkedStack A) -> n-pop {A} {a} (suc n) (pushSingleLinkedStack s a id) ≡ n-pop {A} {a} n s push-and-n-pop zero s = refl push-and-n-pop {A} {a} (suc n) s = begin n-pop (suc (suc n)) (pushSingleLinkedStack s a id) ≡⟨ refl ⟩ popSingleLinkedStack (n-pop (suc n) (pushSingleLinkedStack s a id)) (\s _ -> s) ≡⟨ cong (\s -> popSingleLinkedStack s (\s _ -> s)) (push-and-n-pop n s) ⟩ popSingleLinkedStack (n-pop n s) (\s _ -> s) ≡⟨ refl ⟩ n-pop (suc n) s ∎ n-push-pop-equiv : {A : Set} {a : A} (n : Nat) (s : SingleLinkedStack A) -> (n-pop {A} {a} n (n-push {A} {a} n s)) ≡ s n-push-pop-equiv zero s = refl n-push-pop-equiv {A} {a} (suc n) s = begin n-pop (suc n) (n-push (suc n) s) ≡⟨ refl ⟩ n-pop (suc n) (pushSingleLinkedStack (n-push n s) a (\s -> s)) ≡⟨ push-and-n-pop n (n-push n s) ⟩ n-pop n (n-push n s) ≡⟨ n-push-pop-equiv n s\ ⟩ s ∎