comparison cbc/named-product.agda @ 49:8031568638d0

Define composition of codesegment using subtype without constraint list
author atton <atton@cr.ie.u-ryukyu.ac.jp>
date Fri, 06 Jan 2017 06:29:51 +0000
parents fe5755744071
children
comparison
equal deleted inserted replaced
48:fe5755744071 49:8031568638d0
60 60
61 equiv : incContextCycle firstContext ≡ record { cycle = 4 ; led = "q" ; signature = "aaa" } 61 equiv : incContextCycle firstContext ≡ record { cycle = 4 ; led = "q" ; signature = "aaa" }
62 equiv = refl 62 equiv = refl
63 63
64 64
65 data CodeSegment (A B : Set) : (List Set) -> Set where 65 data CodeSegment (A B : Set) : Set where
66 cs : {l : List Set} {{_ : Datum A}} {{_ : Datum B}} -> (A -> B) -> CodeSegment A B l 66 cs : {{_ : Datum A}} {{_ : Datum B}} -> (A -> B) -> CodeSegment A B
67 67
68 BasicCS : Set -> Set -> Set 68 BasicCS : Set -> Set -> Set
69 BasicCS A B = CodeSegment A B (A ∷ B ∷ []) 69 BasicCS A B = CodeSegment A B
70 70
71 71
72 exec : {I O : Set} -> {l : List Set} {{_ : Datum I}} {{_ : Datum O}} -> CodeSegment I O l -> Context -> Context 72 exec : {I O : Set} {{_ : Datum I}} {{_ : Datum O}} -> CodeSegment I O -> Context -> Context
73 exec {l} {{i}} {{o}} (cs b) c = Datum.set o c (b (get i c)) 73 exec {l} {{i}} {{o}} (cs b) c = Datum.set o c (b (get i c))
74 74
75 75
76 comp : {con : Context} -> {A B C D : Set} {{_ : Datum A}} {{_ : Datum B}} {{_ : Datum C}} {{_ : Datum D}} 76 comp : {con : Context} -> {A B C D : Set} {{_ : Datum A}} {{_ : Datum B}} {{_ : Datum C}} {{_ : Datum D}}
77 -> (C -> D) -> (A -> B) -> A -> D 77 -> (C -> D) -> (A -> B) -> A -> D
85 85
86 equiv-exec : {l : List Set} -> incContextCycle firstContext ≡ exec csInc firstContext 86 equiv-exec : {l : List Set} -> incContextCycle firstContext ≡ exec csInc firstContext
87 equiv-exec = refl 87 equiv-exec = refl
88 88
89 89
90 csComp : {con : Context} {A B C D : Set} {{_ : Datum A}} {{_ : Datum B}} {{_ : Datum C}} {{_ : Datum D}} {l ll : List Set} 90 csComp : {con : Context} {A B C D : Set} {{_ : Datum A}} {{_ : Datum B}} {{_ : Datum C}} {{_ : Datum D}}
91 -> CodeSegment C D l -> CodeSegment A B ll -> CodeSegment A D (l ++ ll) 91 -> CodeSegment C D -> CodeSegment A B -> CodeSegment A D
92 csComp {con} {A} {B} {C} {D} {{da}} {{db}} {{dc}} {{dd}} {l} {ll} (cs g) (cs f) 92 csComp {con} {A} {B} {C} {D} {{da}} {{db}} {{dc}} {{dd}} (cs g) (cs f)
93 = cs {{da}} {{dd}} (comp {con} {{da}} {{db}} {{dc}} {{dd}} g f) 93 = cs {{da}} {{dd}} (comp {con} {{da}} {{db}} {{dc}} {{dd}} g f)
94 94
95 comp-sample : {c : Context} -> CodeSegment LoopCounter LoopCounter (LoopCounter ∷ LoopCounter ∷ LoopCounter ∷ LoopCounter ∷ []) 95 comp-sample : {c : Context} -> CodeSegment LoopCounter LoopCounter
96 comp-sample {c} = (csComp {c} csInc csInc) 96 comp-sample {c} = (csComp {c} csInc csInc)
97 97
98 98
99 apply-sample : Context 99 apply-sample : Context
100 apply-sample = exec (comp-sample {firstContext}) firstContext 100 apply-sample = exec (comp-sample {firstContext}) firstContext
108 csUpdateSignature : BasicCS SignatureWithNum SignatureWithNum 108 csUpdateSignature : BasicCS SignatureWithNum SignatureWithNum
109 csUpdateSignature = cs updateSignature 109 csUpdateSignature = cs updateSignature
110 110
111 111
112 112
113 --comp-sample-2 : {c : Context} -> CodeSegment LoopCounter SignatureWithNum (SignatureWithNum ∷ LoopCounter ∷ SignatureWithNum ∷ LoopCounter ∷ []) 113 comp-sample-2 : {c : Context} -> CodeSegment LoopCounter SignatureWithNum
114 --comp-sample-2 = csComp csUpdateSignature csInc 114 comp-sample-2 {c} = csComp {c} csUpdateSignature csInc
115 115
116 --apply-sample-2 : Context 116 apply-sample-2 : Context
117 --apply-sample-2 = exec (comp-sample-2 {firstContext}) firstContext 117 apply-sample-2 = exec (comp-sample-2 {firstContext}) firstContext
118
119
120 open ≡-Reasoning
121
122 ++empty : {A : Set} -> (l : List A) -> l ++ [] ≡ l
123 ++empty [] = refl
124 ++empty (x ∷ l) = cong (\l -> x ∷ l) (++empty l)
125 118
126 119
127 120
128 list-associative : {A : Set} -> (l ll lll : List A) -> l ++ (ll ++ lll) ≡ (l ++ ll) ++ lll 121 comp-associative : {A B C D E F : Set} {con : Context}
129 list-associative [] ll lll = refl
130 list-associative (x ∷ l) [] [] = begin
131 (x ∷ l) ++ [] ++ []
132 ≡⟨ ++empty ((x ∷ l)) ⟩
133 x ∷ l
134 ≡⟨ sym (++empty (x ∷ l)) ⟩
135 x ∷ l ++ []
136 ≡⟨ sym (++empty ((x ∷ l) ++ [])) ⟩
137 ((x ∷ l) ++ []) ++ []
138
139 list-associative (x ∷ l) [] (xxx ∷ lll) = {!!}
140 list-associative (x ∷ l) (x₁ ∷ ll) [] = {!!}
141 list-associative (x ∷ l) (x₁ ∷ ll) (x₂ ∷ lll) = {!!}
142
143
144
145
146
147 comp-associative : {A B C D E F : Set} {l ll lll : List Set} {con : Context}
148 {{da : Datum A}} {{db : Datum B}} {{dc : Datum C}} {{dd : Datum D}} {{de : Datum E}} {{df : Datum F}} 122 {{da : Datum A}} {{db : Datum B}} {{dc : Datum C}} {{dd : Datum D}} {{de : Datum E}} {{df : Datum F}}
149 -> {a : CodeSegment A B l} {b : CodeSegment C D ll} {c : CodeSegment E F lll} 123 -> (a : CodeSegment A B) (b : CodeSegment C D) (c : CodeSegment E F)
150 -> csComp {con} c (csComp {con} b a) ≡ {!!} -- csComp (csComp c b) a -- cannot define directly 124 -> csComp {con} c (csComp {con} b a) ≡ csComp {con} (csComp {con} c b) a
151 -- csComp c (csComp b a) has CodeSegment A F (lll ++ ll ++ l) 125 comp-associative (cs _) (cs _) (cs _) = refl
152 -- csComp (csComp c b) a has CodeSegment A F ((lll ++ ll) ++ l)
153 comp-associative = {!!}
154