view src/parallel_execution/stack.agda @ 164:b0c6e0392b00

Add comment to stack.agda
author atton
date Thu, 17 Nov 2016 18:28:13 +0000
parents db647f7ed2f6
children bf26f1105862
line wrap: on
line source

module stack where

open import Relation.Binary.PropositionalEquality

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 :  Bool
test03 = pushSingleLinkedStack emptySingleLinkedStack (cons True Nothing) test02 -- ok
--test03 = pushSingleLinkedStack emptySingleLinkedStack (False) test02 -- ok


lemma : test03 ≡ False
lemma = refl