annotate cbc/named-product.agda @ 40:e6b965df2137

Trying define type composition using subtype through DS
author atton <atton@cr.ie.u-ryukyu.ac.jp>
date Tue, 03 Jan 2017 07:52:43 +0000
parents d8312361afdc
children 2abf1cd97f10
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
rev   line source
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
1 module named-product where
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
2
38
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
3 open import Function
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
4 open import Data.Bool
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
5 open import Data.Nat
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
6 open import Data.String
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
7 open import Data.Vec
38
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
8 open import Relation.Binary.PropositionalEquality
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
9
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
10
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
11 data DataSegment {n : ℕ } : (Vec Set n) -> Set where
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
12 ds : (l : Vec Set n) -> DataSegment l
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
13
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
14
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
15 record Context : Set where
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
16 field
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
17 cycle : ℕ
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
18 led : String
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
19 signature : String
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
20
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
21 record Datum (A : Set) : Set where
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
22 field
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
23 get : Context -> A
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
24 set : Context -> A -> Context
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
25
38
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
26
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
27 record LoopCounter : Set where
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
28 field
38
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
29 count : ℕ
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
30 name : String
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
31
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
32 record SignatureWithNum : Set where
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
33 field
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
34 signature : String
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
35 num : ℕ
38
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
36
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
37 instance
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
38 contextHasLoopCounter : Datum LoopCounter
38
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
39 contextHasLoopCounter = record {get = (\c -> record {count = Context.cycle c ;
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
40 name = Context.led c });
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
41 set = (\c l -> record {cycle = LoopCounter.count l;
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
42 led = LoopCounter.name l;
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
43 signature = Context.signature c})}
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
44 contextHasSignatureWithNum : Datum SignatureWithNum
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
45 contextHasSignatureWithNum = record {get = (\c -> record { signature = Context.signature c
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
46 ; num = Context.cycle c})
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
47 ;set = (\c s -> record { cycle = SignatureWithNum.num s
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
48 ; led = Context.led c
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
49 ; signature = SignatureWithNum.signature s})
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
50 }
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
51
38
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
52
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
53 inc : LoopCounter -> LoopCounter
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
54 inc l = record l {count = suc (LoopCounter.count l)}
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
55
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
56 firstContext : Context
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
57 firstContext = record { cycle = 3 ; led = "q" ; signature = "aaa" }
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
58
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
59 incContextCycle : {{_ : Datum LoopCounter }} -> Context -> Context
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
60 incContextCycle {{l}} c = Datum.set l c (inc (Datum.get l c))
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
61
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
62
38
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
63 equiv : incContextCycle firstContext ≡ record { cycle = 4 ; led = "q" ; signature = "aaa" }
a0ca5e29a9dc Call subtype using user-defined cast operator
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 37
diff changeset
64 equiv = refl
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
65
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
66
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
67
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
68 data CodeSegment {n : ℕ} {l : Vec Set n} (A B : Set) : (DataSegment (A ∷ (B ∷ l))) -> Set where
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
69 cs : {{_ : Datum A}} {{_ : Datum B}} -> (A -> B) -> CodeSegment A B (ds (A ∷ (B ∷ l)))
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
70
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
71
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
72 exec : {n : ℕ } {I O : Set} -> {l : Vec Set (suc (suc n))} {{_ : Datum I}} {{_ : Datum O}} -> CodeSegment I O (ds (I ∷ O ∷ l)) -> Context -> Context
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
73 exec {{i}} {{o}} (cs b) c = Datum.set o c (b (Datum.get i c))
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
74
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
75
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
76 equiv-exec : {n : ℕ } {l : Vec Set n} -> incContextCycle firstContext ≡ exec (cs inc) firstContext
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
77 equiv-exec = refl
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
78 {-
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
79
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
80
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
81 InputIsHead : {n : ℕ} {I O : Set} {l : Vec Set (suc (suc n))} -> (cs : CodeSegment I O l) -> I ≡ head l
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
82 InputIsHead (cs _) = refl
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
83
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
84 OutputIsLast : {n : ℕ} {I O : Set} {l : Vec Set (suc (suc n))} -> (cs : CodeSegment I O l) -> O ≡ last l
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
85 OutputIsLast {_} {_} {_} {l} (cs x) = {!!}
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
86
40
e6b965df2137 Trying define type composition using subtype through DS
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 39
diff changeset
87 -}
39
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
88
d8312361afdc Call subtype using user-defined cast operator with subtype list
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents: 38
diff changeset
89
37
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
90
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
91 --yoyo : DataSegment
60e604972f30 Trying define codesegment using named-product with subtype
atton <atton@cr.ie.u-ryukyu.ac.jp>
parents:
diff changeset
92 --yoyo = record { name = "yoyo" ; ds = [ Yo ]}