annotate src/parallel_execution/stack.agda @ 471:e5f0cced7d43

remove error ' rbtree'.
author ryokka
date Wed, 27 Dec 2017 21:17:02 +0900
parents 07ccd411ad70
children c3202635c20a
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
154
efef5d4df54f Add normal level and agda code
atton
parents:
diff changeset
1 module stack where
efef5d4df54f Add normal level and agda code
atton
parents:
diff changeset
2
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
3 open import Relation.Binary.PropositionalEquality
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
4
179
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
5 data Nat : Set where
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
6 zero : Nat
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
7 suc : Nat -> Nat
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
8
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
9 data Bool : Set where
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
10 True : Bool
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
11 False : Bool
164
b0c6e0392b00 Add comment to stack.agda
atton
parents: 161
diff changeset
12
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
13 data Maybe (a : Set) : Set where
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
14 Nothing : Maybe a
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
15 Just : a -> Maybe a
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
16
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
17 record Stack {a t : Set} (stackImpl : Set) : Set where
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
18 field
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
19 stack : stackImpl
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
20 push : stackImpl -> a -> (stackImpl -> t) -> t
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
21 pop : stackImpl -> (stackImpl -> Maybe a -> t) -> t
155
19cc626943e4 stack.agda
atton
parents: 154
diff changeset
22
427
07ccd411ad70 fix interface in agda
ryokka
parents: 181
diff changeset
23 pushStack : {a t : Set} -> Stack -> a -> (Stack -> t) -> t
07ccd411ad70 fix interface in agda
ryokka
parents: 181
diff changeset
24 pushStack {a} {t} s0 d next = (stackImpl s0) (stack s0) d (\s1 -> next (record s0 {stack = s1;} ))
07ccd411ad70 fix interface in agda
ryokka
parents: 181
diff changeset
25
07ccd411ad70 fix interface in agda
ryokka
parents: 181
diff changeset
26 popStack : {a t : Set} -> Stack -> (Stack -> t) -> t
07ccd411ad70 fix interface in agda
ryokka
parents: 181
diff changeset
27 popStack {a} {t} s0 next = (stackImpl s0) (stack s0) (\s1 -> next s0)
07ccd411ad70 fix interface in agda
ryokka
parents: 181
diff changeset
28
07ccd411ad70 fix interface in agda
ryokka
parents: 181
diff changeset
29
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
30 data Element (a : Set) : Set where
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
31 cons : a -> Maybe (Element a) -> Element a
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
32
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
33 datum : {a : Set} -> Element a -> a
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
34 datum (cons a _) = a
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
35
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
36 next : {a : Set} -> Element a -> Maybe (Element a)
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
37 next (cons _ n) = n
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
38
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
39
164
b0c6e0392b00 Add comment to stack.agda
atton
parents: 161
diff changeset
40 {-
b0c6e0392b00 Add comment to stack.agda
atton
parents: 161
diff changeset
41 -- cannot define recrusive record definition. so use linked list with maybe.
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
42 record Element {l : Level} (a : Set l) : Set (suc l) where
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
43 field
164
b0c6e0392b00 Add comment to stack.agda
atton
parents: 161
diff changeset
44 datum : a -- `data` is reserved by Agda.
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
45 next : Maybe (Element a)
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
46 -}
155
19cc626943e4 stack.agda
atton
parents: 154
diff changeset
47
19cc626943e4 stack.agda
atton
parents: 154
diff changeset
48
164
b0c6e0392b00 Add comment to stack.agda
atton
parents: 161
diff changeset
49
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
50 record SingleLinkedStack (a : Set) : Set where
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
51 field
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
52 top : Maybe (Element a)
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
53 open SingleLinkedStack
155
19cc626943e4 stack.agda
atton
parents: 154
diff changeset
54
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
55 pushSingleLinkedStack : {Data t : Set} -> SingleLinkedStack Data -> Data -> (Code : SingleLinkedStack Data -> t) -> t
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
56 pushSingleLinkedStack stack datum next = next stack1
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
57 where
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
58 element = cons datum (top stack)
164
b0c6e0392b00 Add comment to stack.agda
atton
parents: 161
diff changeset
59 stack1 = record {top = Just element}
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
60
155
19cc626943e4 stack.agda
atton
parents: 154
diff changeset
61
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
62 popSingleLinkedStack : {a t : Set} -> SingleLinkedStack a -> (Code : SingleLinkedStack a -> (Maybe a) -> t) -> t
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
63 popSingleLinkedStack stack cs with (top stack)
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
64 ... | Nothing = cs stack Nothing
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
65 ... | Just d = cs stack1 (Just data1)
154
efef5d4df54f Add normal level and agda code
atton
parents:
diff changeset
66 where
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
67 data1 = datum d
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
68 stack1 = record { top = (next d) }
154
efef5d4df54f Add normal level and agda code
atton
parents:
diff changeset
69
efef5d4df54f Add normal level and agda code
atton
parents:
diff changeset
70
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
71 emptySingleLinkedStack : {a : Set} -> SingleLinkedStack a
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
72 emptySingleLinkedStack = record {top = Nothing}
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
73
164
b0c6e0392b00 Add comment to stack.agda
atton
parents: 161
diff changeset
74 createSingleLinkedStack : {a b : Set} -> Stack {a} {b} (SingleLinkedStack a)
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
75 createSingleLinkedStack = record { stack = emptySingleLinkedStack
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
76 ; push = pushSingleLinkedStack
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
77 ; pop = popSingleLinkedStack
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
78 }
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
79
156
0aa2265660a0 Add stack lemma
atton
parents: 155
diff changeset
80
0aa2265660a0 Add stack lemma
atton
parents: 155
diff changeset
81
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
82 test01 : {a : Set} -> SingleLinkedStack a -> Maybe a -> Bool
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
83 test01 stack _ with (top stack)
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
84 ... | (Just _) = True
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
85 ... | Nothing = False
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
86
156
0aa2265660a0 Add stack lemma
atton
parents: 155
diff changeset
87
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
88 test02 : {a : Set} -> SingleLinkedStack a -> Bool
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
89 test02 stack = (popSingleLinkedStack stack) test01
156
0aa2265660a0 Add stack lemma
atton
parents: 155
diff changeset
90
165
bf26f1105862 Generalize lemma
atton
parents: 164
diff changeset
91 test03 : {a : Set} -> a -> Bool
bf26f1105862 Generalize lemma
atton
parents: 164
diff changeset
92 test03 v = pushSingleLinkedStack emptySingleLinkedStack v test02
156
0aa2265660a0 Add stack lemma
atton
parents: 155
diff changeset
93
161
db647f7ed2f6 Prove simple lemma in stack.agda
atton
parents: 158
diff changeset
94
165
bf26f1105862 Generalize lemma
atton
parents: 164
diff changeset
95 lemma : {A : Set} {a : A} -> test03 a ≡ False
158
3fdb3334c7a9 fix stack.cbc
mir3636
parents: 156
diff changeset
96 lemma = refl
179
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
97
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
98 id : {A : Set} -> A -> A
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
99 id a = a
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
100
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
101
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
102 n-push : {A : Set} {a : A} -> Nat -> SingleLinkedStack A -> SingleLinkedStack A
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
103 n-push zero s = s
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
104 n-push {A} {a} (suc n) s = pushSingleLinkedStack (n-push {A} {a} n s) a (\s -> s)
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
105
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
106 n-pop : {A : Set} {a : A} -> Nat -> SingleLinkedStack A -> SingleLinkedStack A
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
107 n-pop zero s = s
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
108 n-pop {A} {a} (suc n) s = popSingleLinkedStack (n-pop {A} {a} n s) (\s _ -> s)
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
109
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
110 open ≡-Reasoning
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
111
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
112 push-pop-equiv : {A : Set} {a : A} (s : SingleLinkedStack A) -> popSingleLinkedStack (pushSingleLinkedStack s a (\s -> s)) (\s _ -> s) ≡ s
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
113 push-pop-equiv s = refl
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
114
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
115 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
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
116 push-and-n-pop zero s = refl
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
117 push-and-n-pop {A} {a} (suc n) s = begin
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
118 n-pop (suc (suc n)) (pushSingleLinkedStack s a id)
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
119 ≡⟨ refl ⟩
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
120 popSingleLinkedStack (n-pop (suc n) (pushSingleLinkedStack s a id)) (\s _ -> s)
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
121 ≡⟨ cong (\s -> popSingleLinkedStack s (\s _ -> s)) (push-and-n-pop n s) ⟩
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
122 popSingleLinkedStack (n-pop n s) (\s _ -> s)
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
123 ≡⟨ refl ⟩
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
124 n-pop (suc n) s
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
125
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
126
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
127
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
128 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
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
129 n-push-pop-equiv zero s = refl
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
130 n-push-pop-equiv {A} {a} (suc n) s = begin
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
131 n-pop (suc n) (n-push (suc n) s)
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
132 ≡⟨ refl ⟩
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
133 n-pop (suc n) (pushSingleLinkedStack (n-push n s) a (\s -> s))
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
134 ≡⟨ push-and-n-pop n (n-push n s) ⟩
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
135 n-pop n (n-push n s)
180
d8947747ff3b Fix syntax
atton
parents: 179
diff changeset
136 ≡⟨ n-push-pop-equiv n s ⟩
179
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
137 s
b3be97ba0782 Prove n-push/n-pop
atton
parents: 165
diff changeset
138
181
78b28c8ffff2 Prove equivalence n-push/n-pop to empty stack
atton
parents: 180
diff changeset
139
78b28c8ffff2 Prove equivalence n-push/n-pop to empty stack
atton
parents: 180
diff changeset
140
78b28c8ffff2 Prove equivalence n-push/n-pop to empty stack
atton
parents: 180
diff changeset
141 n-push-pop-equiv-empty : {A : Set} {a : A} -> (n : Nat) -> n-pop {A} {a} n (n-push {A} {a} n emptySingleLinkedStack) ≡ emptySingleLinkedStack
78b28c8ffff2 Prove equivalence n-push/n-pop to empty stack
atton
parents: 180
diff changeset
142 n-push-pop-equiv-empty n = n-push-pop-equiv n emptySingleLinkedStack