Compare commits
	
		
			No commits in common. "07550bc214138ad995d47c1edd2c6cd6cd0379eb" and "69a4e8eb5c6ab2ae4ba323e2ecd883e4ed98434b" have entirely different histories.
		
	
	
		
			07550bc214
			...
			69a4e8eb5c
		
	
		
| @ -52,7 +52,7 @@ _↑ˡᵉ_ l m = List.map (_↑ˡ m) l | |||||||
| _↑ʳᵉ_ : ∀ {m} n → List (Fin m × Fin m) → List (Fin (n Nat.+ m) × Fin (n Nat.+ m)) | _↑ʳᵉ_ : ∀ {m} n → List (Fin m × Fin m) → List (Fin (n Nat.+ m) × Fin (n Nat.+ m)) | ||||||
| _↑ʳᵉ_ n l = List.map (n ↑ʳ_) l | _↑ʳᵉ_ n l = List.map (n ↑ʳ_) l | ||||||
| 
 | 
 | ||||||
| infixr 5 _∙_ | infixl 5 _∙_ | ||||||
| _∙_ : Graph → Graph → Graph | _∙_ : Graph → Graph → Graph | ||||||
| _∙_ g₁ g₂ = record | _∙_ g₁ g₂ = record | ||||||
|     { size = Graph.size g₁ Nat.+ Graph.size g₂ |     { size = Graph.size g₁ Nat.+ Graph.size g₂ | ||||||
| @ -65,7 +65,7 @@ _∙_ g₁ g₂ = record | |||||||
|                 (Graph.size g₁ ↑ʳⁱ Graph.outputs g₂) |                 (Graph.size g₁ ↑ʳⁱ Graph.outputs g₂) | ||||||
|     } |     } | ||||||
| 
 | 
 | ||||||
| infixr 5 _↦_ | infixl 5 _↦_ | ||||||
| _↦_ : Graph → Graph → Graph | _↦_ : Graph → Graph → Graph | ||||||
| _↦_ g₁ g₂ = record | _↦_ g₁ g₂ = record | ||||||
|     { size = Graph.size g₁ Nat.+ Graph.size g₂ |     { size = Graph.size g₁ Nat.+ Graph.size g₂ | ||||||
| @ -84,7 +84,7 @@ loop g = record g | |||||||
|               List.cartesianProduct (Graph.outputs g) (Graph.inputs g) |               List.cartesianProduct (Graph.outputs g) (Graph.inputs g) | ||||||
|     } |     } | ||||||
| 
 | 
 | ||||||
| infixr 5 _skipto_ | infixl 5 _skipto_ | ||||||
| _skipto_ : Graph → Graph → Graph | _skipto_ : Graph → Graph → Graph | ||||||
| _skipto_ g₁ g₂ = record (g₁ ∙ g₂) | _skipto_ g₁ g₂ = record (g₁ ∙ g₂) | ||||||
|     { edges = Graph.edges (g₁ ∙ g₂) List.++ |     { edges = Graph.edges (g₁ ∙ g₂) List.++ | ||||||
|  | |||||||
| @ -32,87 +32,57 @@ buildCfg-output (if _ then s₁ else s₂) = (_ , refl) | |||||||
| buildCfg-output (while _ repeat s) | buildCfg-output (while _ repeat s) | ||||||
|     with (idx , p) ← buildCfg-output s rewrite p = (_ , refl) |     with (idx , p) ← buildCfg-output s rewrite p = (_ , refl) | ||||||
| 
 | 
 | ||||||
| Trace-∙ˡ : ∀ {g₁ g₂ : Graph} {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env} → | Trace-∙ˡ : ∀ (g₁ g₂ : Graph) {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env} → | ||||||
|            Trace {g₁} idx₁ idx₂ ρ₁ ρ₂ → |            Trace {g₁} idx₁ idx₂ ρ₁ ρ₂ → | ||||||
|            Trace {g₁ ∙ g₂} (idx₁ Fin.↑ˡ Graph.size g₂) (idx₂ Fin.↑ˡ Graph.size g₂) ρ₁ ρ₂ |            Trace {g₁ ∙ g₂} (idx₁ Fin.↑ˡ Graph.size g₂) (idx₂ Fin.↑ˡ Graph.size g₂) ρ₁ ρ₂ | ||||||
| Trace-∙ˡ {g₁} {g₂} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) | Trace-∙ˡ g₁ g₂ {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) | ||||||
|     rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = |     rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = | ||||||
|     Trace-single ρ₁⇒ρ₂ |     Trace-single ρ₁⇒ρ₂ | ||||||
| Trace-∙ˡ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') | Trace-∙ˡ g₁ g₂ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') | ||||||
|     rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = |     rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = | ||||||
|     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (_↑ˡ Graph.size g₂) idx₁→idx)) |     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (_↑ˡ Graph.size g₂) idx₁→idx)) | ||||||
|                     (Trace-∙ˡ tr') |                     (Trace-∙ˡ g₁ g₂ tr') | ||||||
| 
 | 
 | ||||||
| Trace-∙ʳ : ∀ {g₁ g₂ : Graph} {idx₁ idx₂ : Graph.Index g₂} {ρ₁ ρ₂ : Env} → | Trace-∙ʳ : ∀ (g₁ g₂ : Graph) {idx₁ idx₂ : Graph.Index g₂} {ρ₁ ρ₂ : Env} → | ||||||
|            Trace {g₂} idx₁ idx₂ ρ₁ ρ₂ → |            Trace {g₂} idx₁ idx₂ ρ₁ ρ₂ → | ||||||
|            Trace {g₁ ∙ g₂} (Graph.size g₁ Fin.↑ʳ idx₁) (Graph.size g₁ Fin.↑ʳ idx₂) ρ₁ ρ₂ |            Trace {g₁ ∙ g₂} (Graph.size g₁ Fin.↑ʳ idx₁) (Graph.size g₁ Fin.↑ʳ idx₂) ρ₁ ρ₂ | ||||||
| Trace-∙ʳ {g₁} {g₂} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) | Trace-∙ʳ g₁ g₂ {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) | ||||||
|     rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = |     rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = | ||||||
|     Trace-single ρ₁⇒ρ₂ |     Trace-single ρ₁⇒ρ₂ | ||||||
| Trace-∙ʳ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') | Trace-∙ʳ g₁ g₂ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') | ||||||
|     rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = |     rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = | ||||||
|     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ʳ _ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx)) |     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ʳ _ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx)) | ||||||
|                     (Trace-∙ʳ tr') |                     (Trace-∙ʳ g₁ g₂ tr') | ||||||
| 
 | 
 | ||||||
| EndToEndTrace-∙ˡ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ : Env} → | Trace-↦ˡ : ∀ (g₁ g₂ : Graph) {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env} → | ||||||
|                    EndToEndTrace {g₁} ρ₁ ρ₂ → |  | ||||||
|                    EndToEndTrace {g₁ ∙ g₂} ρ₁ ρ₂ |  | ||||||
| EndToEndTrace-∙ˡ {g₁} {g₂} etr = record |  | ||||||
|     { idx₁ = EndToEndTrace.idx₁ etr Fin.↑ˡ Graph.size g₂ |  | ||||||
|     ; idx₁∈inputs = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂) |  | ||||||
|                                                     (EndToEndTrace.idx₁∈inputs etr)) |  | ||||||
|     ; idx₂ = EndToEndTrace.idx₂ etr Fin.↑ˡ Graph.size g₂ |  | ||||||
|     ; idx₂∈outputs = ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Fin._↑ˡ Graph.size g₂) |  | ||||||
|                                                     (EndToEndTrace.idx₂∈outputs etr)) |  | ||||||
|     ; trace = Trace-∙ˡ (EndToEndTrace.trace etr) |  | ||||||
|     } |  | ||||||
| 
 |  | ||||||
| EndToEndTrace-∙ʳ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ : Env} → |  | ||||||
|            EndToEndTrace {g₂} ρ₁ ρ₂ → |  | ||||||
|            EndToEndTrace {g₁ ∙ g₂} ρ₁ ρ₂ |  | ||||||
| EndToEndTrace-∙ʳ {g₁} {g₂} etr = record |  | ||||||
|     { idx₁ = Graph.size g₁ Fin.↑ʳ EndToEndTrace.idx₁ etr |  | ||||||
|     ; idx₁∈inputs = ListMemProp.∈-++⁺ʳ (Graph.inputs g₁ ↑ˡⁱ Graph.size g₂) |  | ||||||
|                                        ((x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_) |  | ||||||
|                                                      (EndToEndTrace.idx₁∈inputs etr))) |  | ||||||
|     ; idx₂ = Graph.size g₁ Fin.↑ʳ EndToEndTrace.idx₂ etr |  | ||||||
|     ; idx₂∈outputs = ListMemProp.∈-++⁺ʳ (Graph.outputs g₁ ↑ˡⁱ Graph.size g₂) |  | ||||||
|                                         ((x∈xs⇒fx∈fxs (Graph.size g₁ Fin.↑ʳ_) |  | ||||||
|                                                       (EndToEndTrace.idx₂∈outputs etr))) |  | ||||||
| 
 |  | ||||||
|     ; trace = Trace-∙ʳ (EndToEndTrace.trace etr) |  | ||||||
|     } |  | ||||||
| 
 |  | ||||||
| Trace-↦ˡ : ∀ {g₁ g₂ : Graph} {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env} → |  | ||||||
|            Trace {g₁} idx₁ idx₂ ρ₁ ρ₂ → |            Trace {g₁} idx₁ idx₂ ρ₁ ρ₂ → | ||||||
|            Trace {g₁ ↦ g₂} (idx₁ Fin.↑ˡ Graph.size g₂) (idx₂ Fin.↑ˡ Graph.size g₂) ρ₁ ρ₂ |            Trace {g₁ ↦ g₂} (idx₁ Fin.↑ˡ Graph.size g₂) (idx₂ Fin.↑ˡ Graph.size g₂) ρ₁ ρ₂ | ||||||
| Trace-↦ˡ {g₁} {g₂} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) | Trace-↦ˡ g₁ g₂ {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) | ||||||
|     rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = |     rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = | ||||||
|     Trace-single ρ₁⇒ρ₂ |     Trace-single ρ₁⇒ρ₂ | ||||||
| Trace-↦ˡ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') | Trace-↦ˡ g₁ g₂ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') | ||||||
|     rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = |     rewrite sym (lookup-++ˡ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = | ||||||
|     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (_↑ˡ Graph.size g₂) idx₁→idx)) |     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (_↑ˡ Graph.size g₂) idx₁→idx)) | ||||||
|                     (Trace-↦ˡ tr') |                     (Trace-↦ˡ g₁ g₂ tr') | ||||||
| 
 | 
 | ||||||
| Trace-↦ʳ : ∀ {g₁ g₂ : Graph} {idx₁ idx₂ : Graph.Index g₂} {ρ₁ ρ₂ : Env} → | Trace-↦ʳ : ∀ (g₁ g₂ : Graph) {idx₁ idx₂ : Graph.Index g₂} {ρ₁ ρ₂ : Env} → | ||||||
|            Trace {g₂} idx₁ idx₂ ρ₁ ρ₂ → |            Trace {g₂} idx₁ idx₂ ρ₁ ρ₂ → | ||||||
|            Trace {g₁ ↦ g₂} (Graph.size g₁ Fin.↑ʳ idx₁) (Graph.size g₁ Fin.↑ʳ idx₂) ρ₁ ρ₂ |            Trace {g₁ ↦ g₂} (Graph.size g₁ Fin.↑ʳ idx₁) (Graph.size g₁ Fin.↑ʳ idx₂) ρ₁ ρ₂ | ||||||
| Trace-↦ʳ {g₁} {g₂} {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) | Trace-↦ʳ g₁ g₂ {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) | ||||||
|     rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = |     rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = | ||||||
|     Trace-single ρ₁⇒ρ₂ |     Trace-single ρ₁⇒ρ₂ | ||||||
| Trace-↦ʳ {g₁} {g₂} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') | Trace-↦ʳ g₁ g₂ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') | ||||||
|     rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = |     rewrite sym (lookup-++ʳ (Graph.nodes g₁) (Graph.nodes g₂) idx₁) = | ||||||
|     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ʳ (Graph.edges g₁ ↑ˡᵉ Graph.size g₂) |     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ʳ (Graph.edges g₁ ↑ˡᵉ Graph.size g₂) | ||||||
|                                         (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx))) |                                         (ListMemProp.∈-++⁺ˡ (x∈xs⇒fx∈fxs (Graph.size g₁ ↑ʳ_) idx₁→idx))) | ||||||
|                     (Trace-↦ʳ {g₁} {g₂} tr') |                     (Trace-↦ʳ g₁ g₂ tr') | ||||||
| 
 | 
 | ||||||
| Trace-loop : ∀ {g₁ : Graph} {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env} → | Trace-loop : ∀ (g₁ : Graph) {idx₁ idx₂ : Graph.Index g₁} {ρ₁ ρ₂ : Env} → | ||||||
|              Trace {g₁} idx₁ idx₂ ρ₁ ρ₂ → Trace {loop g₁} idx₁ idx₂ ρ₁ ρ₂ |              Trace {g₁} idx₁ idx₂ ρ₁ ρ₂ → Trace {loop g₁} idx₁ idx₂ ρ₁ ρ₂ | ||||||
| Trace-loop {idx₁ = idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) = Trace-single ρ₁⇒ρ₂ | Trace-loop g₁ {idx₁} {idx₁} (Trace-single ρ₁⇒ρ₂) = Trace-single ρ₁⇒ρ₂ | ||||||
| Trace-loop {g₁} {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') = | Trace-loop g₁ {idx₁} (Trace-edge ρ₁⇒ρ idx₁→idx tr') = | ||||||
|     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ idx₁→idx) (Trace-loop tr') |     Trace-edge ρ₁⇒ρ (ListMemProp.∈-++⁺ˡ idx₁→idx) (Trace-loop g₁ tr') | ||||||
| 
 | 
 | ||||||
| infixr 5 _++_ |  | ||||||
| _++_ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ ρ₃ : Env} → | _++_ : ∀ {g₁ g₂ : Graph} {ρ₁ ρ₂ ρ₃ : Env} → | ||||||
|        EndToEndTrace {g₁} ρ₁ ρ₂ → EndToEndTrace {g₂} ρ₂ ρ₃ → |        EndToEndTrace {g₁} ρ₁ ρ₂ → EndToEndTrace {g₂} ρ₂ ρ₃ → | ||||||
|        EndToEndTrace {g₁ ↦ g₂} ρ₁ ρ₃ |        EndToEndTrace {g₁ ↦ g₂} ρ₁ ρ₃ | ||||||
| @ -130,8 +100,8 @@ _++_ {g₁} {g₂} etr₁ etr₂ | |||||||
|                                            (ListMemProp.∈-++⁺ʳ (Graph.size g₁ ↑ʳᵉ Graph.edges g₂) |                                            (ListMemProp.∈-++⁺ʳ (Graph.size g₁ ↑ʳᵉ Graph.edges g₂) | ||||||
|                                                                (∈-cartesianProduct o∈tr₁ i∈tr₂)) |                                                                (∈-cartesianProduct o∈tr₁ i∈tr₂)) | ||||||
|             in |             in | ||||||
|                 (Trace-↦ˡ {g₁} {g₂} (EndToEndTrace.trace etr₁)) ++⟨ oi∈es ⟩ |                 (Trace-↦ˡ g₁ g₂ (EndToEndTrace.trace etr₁)) ++⟨ oi∈es ⟩ | ||||||
|                 (Trace-↦ʳ {g₁} {g₂} (EndToEndTrace.trace etr₂)) |                 (Trace-↦ʳ g₁ g₂ (EndToEndTrace.trace etr₂)) | ||||||
|         } |         } | ||||||
| 
 | 
 | ||||||
| Trace-singleton : ∀ {bss : List BasicStmt} {ρ₁ ρ₂ : Env} → | Trace-singleton : ∀ {bss : List BasicStmt} {ρ₁ ρ₂ : Env} → | ||||||
| @ -153,11 +123,3 @@ buildCfg-sufficient (⇒ˢ-⟨⟩ ρ₁ ρ₂ bs ρ₁,bs⇒ρ₂) = | |||||||
|     Trace-singleton (ρ₁,bs⇒ρ₂ ∷ []) |     Trace-singleton (ρ₁,bs⇒ρ₂ ∷ []) | ||||||
| buildCfg-sufficient (⇒ˢ-then ρ₁ ρ₂ ρ₃ s₁ s₂ ρ₁,s₁⇒ρ₂ ρ₂,s₂⇒ρ₃) = | buildCfg-sufficient (⇒ˢ-then ρ₁ ρ₂ ρ₃ s₁ s₂ ρ₁,s₁⇒ρ₂ ρ₂,s₂⇒ρ₃) = | ||||||
|     buildCfg-sufficient ρ₁,s₁⇒ρ₂ ++ buildCfg-sufficient ρ₂,s₂⇒ρ₃ |     buildCfg-sufficient ρ₁,s₁⇒ρ₂ ++ buildCfg-sufficient ρ₂,s₂⇒ρ₃ | ||||||
| buildCfg-sufficient (⇒ˢ-if-true ρ₁ ρ₂ _ _ s₁ s₂ _ _ ρ₁,s₁⇒ρ₂) = |  | ||||||
|     Trace-singleton[] ρ₁ ++ |  | ||||||
|     (EndToEndTrace-∙ˡ (buildCfg-sufficient ρ₁,s₁⇒ρ₂)) ++ |  | ||||||
|     Trace-singleton[] ρ₂ |  | ||||||
| buildCfg-sufficient (⇒ˢ-if-false ρ₁ ρ₂ _ s₁ s₂ _ ρ₁,s₂⇒ρ₂) = |  | ||||||
|     Trace-singleton[] ρ₁ ++ |  | ||||||
|     (EndToEndTrace-∙ʳ {buildCfg s₁} (buildCfg-sufficient ρ₁,s₂⇒ρ₂)) ++ |  | ||||||
|     Trace-singleton[] ρ₂ |  | ||||||
|  | |||||||
		Loading…
	
		Reference in New Issue
	
	Block a user