Mercurial > hg > Members > kono > Proof > automaton
annotate automaton-in-agda/src/non-regular.agda @ 395:cd81e0869fac
...
author | Shinji KONO <kono@ie.u-ryukyu.ac.jp> |
---|---|
date | Thu, 03 Aug 2023 14:55:14 +0900 |
parents | d860e357fe5f |
children | f26444eb0275 |
rev | line source |
---|---|
141 | 1 module non-regular where |
2 | |
3 open import Data.Nat | |
274 | 4 open import Data.Empty |
141 | 5 open import Data.List |
278
e89957b99662
dup in finiteSet in long list
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
277
diff
changeset
|
6 open import Data.Maybe hiding ( map ) |
141 | 7 open import Relation.Binary.PropositionalEquality hiding ( [_] ) |
8 open import logic | |
9 open import automaton | |
274 | 10 open import automaton-ex |
278
e89957b99662
dup in finiteSet in long list
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
277
diff
changeset
|
11 open import finiteSetUtil |
141 | 12 open import finiteSet |
392 | 13 open import Relation.Nullary |
274 | 14 open import regular-language |
306 | 15 open import nat |
385 | 16 open import pumping |
306 | 17 |
141 | 18 |
274 | 19 open FiniteSet |
20 | |
385 | 21 list-eq : List In2 → List In2 → Bool |
22 list-eq [] [] = true | |
23 list-eq [] (x ∷ s1) = false | |
24 list-eq (x ∷ s) [] = false | |
25 list-eq (i0 ∷ s) (i0 ∷ s1) = false | |
26 list-eq (i0 ∷ s) (i1 ∷ s1) = false | |
27 list-eq (i1 ∷ s) (i0 ∷ s1) = false | |
392 | 28 list-eq (i1 ∷ s) (i1 ∷ s1) = list-eq s s1 |
385 | 29 |
392 | 30 input-addi0 : ( n : ℕ ) → List In2 → List In2 |
385 | 31 input-addi0 zero x = x |
32 input-addi0 (suc i) x = i0 ∷ input-addi0 i x | |
33 | |
392 | 34 input-addi1 : ( n : ℕ ) → List In2 |
385 | 35 input-addi1 zero = [] |
36 input-addi1 (suc i) = i1 ∷ input-addi1 i | |
274 | 37 |
392 | 38 inputnn0 : ( n : ℕ ) → List In2 |
385 | 39 inputnn0 n = input-addi0 n (input-addi1 n) |
40 | |
41 inputnn1-i1 : (i : ℕ) → List In2 → Bool | |
42 inputnn1-i1 zero [] = true | |
43 inputnn1-i1 (suc _) [] = false | |
44 inputnn1-i1 zero (i1 ∷ x) = false | |
45 inputnn1-i1 (suc i) (i1 ∷ x) = inputnn1-i1 i x | |
46 inputnn1-i1 zero (i0 ∷ x) = false | |
47 inputnn1-i1 (suc _) (i0 ∷ x) = false | |
48 | |
49 inputnn1-i0 : (i : ℕ) → List In2 → ℕ ∧ List In2 | |
50 inputnn1-i0 i [] = ⟪ i , [] ⟫ | |
51 inputnn1-i0 i (i1 ∷ x) = ⟪ i , (i1 ∷ x) ⟫ | |
392 | 52 inputnn1-i0 i (i0 ∷ x) = inputnn1-i0 (suc i) x |
385 | 53 |
54 open _∧_ | |
55 | |
56 inputnn1 : List In2 → Bool | |
57 inputnn1 x = inputnn1-i1 (proj1 (inputnn1-i0 0 x)) (proj2 (inputnn1-i0 0 x)) | |
274 | 58 |
59 t1 = inputnn1 ( i0 ∷ i1 ∷ [] ) | |
60 t2 = inputnn1 ( i0 ∷ i0 ∷ i1 ∷ i1 ∷ [] ) | |
277 | 61 t3 = inputnn1 ( i0 ∷ i0 ∷ i0 ∷ i1 ∷ i1 ∷ [] ) |
274 | 62 |
385 | 63 t4 : inputnn1 ( inputnn0 5 ) ≡ true |
274 | 64 t4 = refl |
65 | |
291 | 66 t5 : ( n : ℕ ) → Set |
385 | 67 t5 n = inputnn1 ( inputnn0 n ) ≡ true |
68 | |
69 cons-inject : {A : Set} {x1 x2 : List A } { a : A } → a ∷ x1 ≡ a ∷ x2 → x1 ≡ x2 | |
70 cons-inject refl = refl | |
71 | |
72 append-[] : {A : Set} {x1 : List A } → x1 ++ [] ≡ x1 | |
73 append-[] {A} {[]} = refl | |
74 append-[] {A} {x ∷ x1} = cong (λ k → x ∷ k) (append-[] {A} {x1} ) | |
75 | |
76 open import Data.Nat.Properties | |
77 open import Relation.Binary.Definitions | |
78 open import Relation.Binary.PropositionalEquality | |
291 | 79 |
388 | 80 nn30 : (y : List In2) → (j : ℕ) → proj2 (inputnn1-i0 (suc j) y) ≡ proj2 (inputnn1-i0 j y ) |
81 nn30 [] _ = refl | |
82 nn30 (i0 ∷ y) j = nn30 y (suc j) | |
83 nn30 (i1 ∷ y) _ = refl | |
84 | |
85 nn31 : (y : List In2) → (j : ℕ) → proj1 (inputnn1-i0 (suc j) y) ≡ suc (proj1 (inputnn1-i0 j y )) | |
86 nn31 [] _ = refl | |
87 nn31 (i0 ∷ y) j = nn31 y (suc j) | |
88 nn31 (i1 ∷ y) _ = refl | |
89 | |
385 | 90 nn01 : (i : ℕ) → inputnn1 ( inputnn0 i ) ≡ true |
91 nn01 i = subst₂ (λ j k → inputnn1-i1 j k ≡ true) (sym (nn07 i 0 refl)) (sym (nn09 i)) (nn04 i) where | |
92 nn07 : (j x : ℕ) → x + j ≡ i → proj1 ( inputnn1-i0 x (input-addi0 j (input-addi1 i))) ≡ x + j | |
93 nn07 zero x eq with input-addi1 i | inspect input-addi1 i | |
94 ... | [] | _ = +-comm 0 _ | |
95 ... | i0 ∷ t | record { eq = eq1 } = ⊥-elim ( nn08 i eq1 ) where | |
96 nn08 : (i : ℕ) → ¬ (input-addi1 i ≡ i0 ∷ t ) | |
97 nn08 zero () | |
98 nn08 (suc i) () | |
99 ... | i1 ∷ t | _ = +-comm 0 _ | |
392 | 100 nn07 (suc j) x eq = trans (nn07 j (suc x) (trans (cong (λ k → k + j) (+-comm 1 _ )) (trans (+-assoc x _ _) eq)) ) |
385 | 101 (trans (+-assoc 1 x _) (trans (cong (λ k → k + j) (+-comm 1 _) ) (+-assoc x 1 _) )) |
102 nn09 : (x : ℕ) → proj2 ( inputnn1-i0 0 (input-addi0 x (input-addi1 i))) ≡ input-addi1 i | |
103 nn09 zero with input-addi1 i | inspect input-addi1 i | |
104 ... | [] | _ = refl | |
105 ... | i0 ∷ t | record { eq = eq1 } = ⊥-elim ( nn08 i eq1 ) where | |
106 nn08 : (i : ℕ) → ¬ (input-addi1 i ≡ i0 ∷ t ) | |
107 nn08 zero () | |
108 nn08 (suc i) () | |
109 ... | i1 ∷ t | _ = refl | |
392 | 110 nn09 (suc j) = trans (nn30 (input-addi0 j (input-addi1 i)) 0) (nn09 j ) |
385 | 111 nn04 : (i : ℕ) → inputnn1-i1 i (input-addi1 i) ≡ true |
112 nn04 zero = refl | |
113 nn04 (suc i) = nn04 i | |
114 | |
393 | 115 half : (x : List In2) → ℕ |
116 half [] = 0 | |
117 half (x ∷ []) = 0 | |
118 half (x ∷ x₁ ∷ x₂) = suc (half x₂) | |
119 | |
120 top-is-i0 : (x : List In2) → Bool | |
121 top-is-i0 [] = true | |
122 top-is-i0 (i0 ∷ _) = true | |
123 top-is-i0 (i1 ∷ _) = false | |
124 | |
125 nn02 : (x : List In2) → inputnn1 x ≡ true → x ≡ inputnn0 (half x) | |
395 | 126 nn02 x eq = nn08 x eq where |
394 | 127 nn07 : (i : ℕ) → (x : List In2) → inputnn1-i1 i x ≡ true → x ≡ input-addi1 i |
128 nn07 zero [] eq = refl | |
129 nn07 zero (i0 ∷ x₁) () | |
130 nn07 zero (i1 ∷ x₁) () | |
131 nn07 (suc i) (i1 ∷ x₁) eq = cong (λ k → i1 ∷ k) (nn07 i x₁ eq) | |
395 | 132 nn08 : (x : List In2 ) → inputnn1-i1 (proj1 (inputnn1-i0 0 x)) (proj2 (inputnn1-i0 0 x)) ≡ true |
133 → x ≡ input-addi0 (half x) (input-addi1 (half x)) | |
134 nn08 [] eq = ? | |
135 nn08 (i1 ∷ t) eq = ? | |
136 nn08 (i0 ∷ []) eq = ? | |
137 nn08 (i0 ∷ i1 ∷ t) eq = ? | |
138 nn08 (i0 ∷ t @ (i0 ∷ _)) eq = ? where | |
139 nn20 : top-is-i0 t ≡ true | |
140 nn20 = refl | |
141 y : List In2 → List In2 | |
142 y [] = [] | |
143 y (x ∷ []) = [] | |
144 y (x ∷ t ∷ z ) = x ∷ y (t ∷ z) | |
145 nn15 : ( x y : List In2) → length x ≡ suc (suc (length y)) → half x ≡ suc (half y) | |
146 nn15 (x ∷ x₁ ∷ []) [] eq = refl | |
147 nn15 (x ∷ x₁ ∷ x₃ ∷ []) (x₂ ∷ []) eq = refl | |
148 nn15 (x ∷ x₁ ∷ x₃ ∷ x₅ ∷ []) (x₂ ∷ x₄ ∷ []) eq = refl | |
149 nn15 (x ∷ x₁ ∷ x₃ ∷ x₅) (x₂ ∷ x₄ ∷ x₆ ∷ y₁) eq = cong suc (nn15 (x₃ ∷ x₅) (x₆ ∷ y₁) (cong pred (cong pred eq)) ) | |
150 nn13 : (z : List In2) → half (i0 ∷ i0 ∷ z) ≡ suc (half (y (i0 ∷ z))) | |
151 nn13 z = nn15 (i0 ∷ i0 ∷ z) (y (i0 ∷ z)) ? where | |
152 nn17 : {x : In2} (z : List In2) → length (y (i0 ∷ z)) ≡ length (y (x ∷ z)) | |
153 nn17 [] = refl | |
154 nn17 (x ∷ []) = refl | |
155 nn17 (x ∷ x₁ ∷ z) = cong suc ( nn17 {x} (x₁ ∷ z)) | |
156 nn16 : (z : List In2 ) → length (i0 ∷ i0 ∷ z) ≡ suc (suc (length (y (i0 ∷ z)))) | |
157 nn16 [] = refl | |
158 nn16 (x ∷ z) = begin | |
159 length (i0 ∷ i0 ∷ x ∷ z) ≡⟨ refl ⟩ | |
160 suc (length (i0 ∷ i0 ∷ z)) ≡⟨ cong suc (nn16 z) ⟩ | |
161 suc (suc (suc (length (y (i0 ∷ z))))) ≡⟨ cong suc (cong suc ( cong suc (nn17 z))) ⟩ | |
162 suc (suc (suc (length (y (x ∷ z))))) ≡⟨ refl ⟩ | |
163 suc (suc (length (i0 ∷ y (x ∷ z)))) ∎ where open ≡-Reasoning | |
164 nn11 : (t : List In2) → inputnn1-i1 (proj1 (inputnn1-i0 1 t)) (proj2 (inputnn1-i0 1 t)) ≡ true | |
165 → t ++ (i1 ∷ []) ≡ y t | |
166 nn11 t eq = ? | |
167 nn10 : (y : List In2 ) | |
168 → inputnn1-i1 (proj1 (inputnn1-i0 0 y)) (proj2 (inputnn1-i0 0 y)) ≡ true | |
169 → y ≡ input-addi0 (half y) (input-addi1 (half y)) | |
170 nn10 y eq = nn08 y eq | |
171 nn14 : inputnn1-i1 (proj1 (inputnn1-i0 0 (y t))) (proj2 (inputnn1-i0 0 (y t))) ≡ true | |
172 nn14 = ? | |
173 nn12 : i0 ∷ t ≡ input-addi0 (half (i0 ∷ t)) (input-addi1 (half (i0 ∷ t))) | |
174 nn12 = ? | |
175 | |
176 | |
274 | 177 -- |
178 -- if there is an automaton with n states , which accespt inputnn1, it has a trasition function. | |
179 -- The function is determinted by inputs, | |
180 -- | |
181 | |
182 open RegularLanguage | |
183 open Automaton | |
184 | |
185 open _∧_ | |
141 | 186 |
277 | 187 |
280 | 188 open RegularLanguage |
294 | 189 open import Data.Nat.Properties |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
190 open import Data.List.Properties |
294 | 191 open import nat |
280 | 192 |
392 | 193 lemmaNN : (r : RegularLanguage In2 ) → ¬ ( (s : List In2) → isRegular inputnn1 s r ) |
332 | 194 lemmaNN r Rg = tann {TA.x TAnn} (TA.non-nil-y TAnn ) (TA.xyz=is TAnn) (tr-accept→ (automaton r) _ (astart r) (TA.trace-xyz TAnn) ) |
317 | 195 (tr-accept→ (automaton r) _ (astart r) (TA.trace-xyyz TAnn) ) where |
280 | 196 n : ℕ |
197 n = suc (finite (afin r)) | |
392 | 198 nn = inputnn0 n |
280 | 199 nn03 : accept (automaton r) (astart r) nn ≡ true |
294 | 200 nn03 = subst (λ k → k ≡ true ) (Rg nn ) (nn01 n) |
304 | 201 nn09 : (n m : ℕ) → n ≤ n + m |
202 nn09 zero m = z≤n | |
203 nn09 (suc n) m = s≤s (nn09 n m) | |
295 | 204 nn04 : Trace (automaton r) nn (astart r) |
392 | 205 nn04 = tr-accept← (automaton r) nn (astart r) nn03 |
315 | 206 nntrace = tr→qs (automaton r) nn (astart r) nn04 |
392 | 207 nn07 : (n : ℕ) → length (inputnn0 n ) ≡ n + n |
385 | 208 nn07 i = nn19 i where |
209 nn17 : (i : ℕ) → length (input-addi1 i) ≡ i | |
210 nn17 zero = refl | |
211 nn17 (suc i)= cong suc (nn17 i) | |
212 nn18 : (i j : ℕ) → length (input-addi0 j (input-addi1 i)) ≡ j + length (input-addi1 i ) | |
213 nn18 i zero = refl | |
214 nn18 i (suc j)= cong suc (nn18 i j) | |
215 nn19 : (i : ℕ) → length (input-addi0 i ( input-addi1 i )) ≡ i + i | |
216 nn19 i = begin | |
217 length (input-addi0 i ( input-addi1 i )) ≡⟨ nn18 i i ⟩ | |
218 i + length (input-addi1 i) ≡⟨ cong (λ k → i + k) ( nn17 i) ⟩ | |
392 | 219 i + i ∎ where open ≡-Reasoning |
294 | 220 nn05 : length nntrace > finite (afin r) |
221 nn05 = begin | |
222 suc (finite (afin r)) ≤⟨ nn09 _ _ ⟩ | |
223 n + n ≡⟨ sym (nn07 n) ⟩ | |
385 | 224 length (inputnn0 n ) ≡⟨ tr→qs=is (automaton r) (inputnn0 n ) (astart r) nn04 ⟩ |
294 | 225 length nntrace ∎ where open ≤-Reasoning |
315 | 226 nn06 : Dup-in-list ( afin r) (tr→qs (automaton r) nn (astart r) nn04) |
304 | 227 nn06 = dup-in-list>n (afin r) nntrace nn05 |
332 | 228 |
304 | 229 TAnn : TA (automaton r) (astart r) nn |
317 | 230 TAnn = pumping-lemma (automaton r) (afin r) (astart r) (Dup-in-list.dup nn06) nn nn04 (Dup-in-list.is-dup nn06) |
332 | 231 |
385 | 232 open import Tactic.MonoidSolver using (solve; solve-macro) |
233 | |
392 | 234 -- there is a counter example |
235 -- | |
317 | 236 tann : {x y z : List In2} → ¬ y ≡ [] |
237 → x ++ y ++ z ≡ nn | |
238 → accept (automaton r) (astart r) (x ++ y ++ z) ≡ true → ¬ (accept (automaton r) (astart r) (x ++ y ++ y ++ z) ≡ true ) | |
387 | 239 tann {x} {y} {z} ny eq axyz axyyz = ¬-bool (nn10 x y z ny (trans (Rg (x ++ y ++ z)) axyz ) ) (trans (Rg (x ++ y ++ y ++ z)) axyyz ) where |
385 | 240 count0 : (x : List In2) → ℕ |
241 count0 [] = 0 | |
242 count0 (i0 ∷ x) = suc (count0 x) | |
243 count0 (i1 ∷ x) = count0 x | |
244 count1 : (x : List In2) → ℕ | |
245 count1 [] = 0 | |
246 count1 (i1 ∷ x) = suc (count1 x) | |
247 count1 (i0 ∷ x) = count1 x | |
392 | 248 -- |
249 -- prove some obvious fact | |
250 -- | |
387 | 251 c0+1=n : (x : List In2 ) → count0 x + count1 x ≡ length x |
252 c0+1=n [] = refl | |
253 c0+1=n (i0 ∷ t) = cong suc ( c0+1=n t ) | |
254 c0+1=n (i1 ∷ t) = begin | |
392 | 255 count0 t + suc (count1 t) ≡⟨ sym (+-assoc (count0 t) _ _) ⟩ |
256 (count0 t + 1 ) + count1 t ≡⟨ cong (λ k → k + count1 t) (+-comm _ 1 ) ⟩ | |
257 suc (count0 t + count1 t) ≡⟨ cong suc ( c0+1=n t ) ⟩ | |
387 | 258 suc (length t) ∎ where open ≡-Reasoning |
392 | 259 -- |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
260 nn15 : (x : List In2 ) → inputnn1 x ≡ true → count0 x ≡ count1 x |
388 | 261 nn15 x eq = nn18 where |
392 | 262 nn17 : (x : List In2 ) → (count0 x ≡ proj1 (inputnn1-i0 0 x) + count0 (proj2 (inputnn1-i0 0 x))) |
388 | 263 ∧ (count1 x ≡ 0 + count1 (proj2 (inputnn1-i0 0 x))) |
264 nn17 [] = ⟪ refl , refl ⟫ | |
265 nn17 (i0 ∷ t ) with nn17 t | |
266 ... | ⟪ eq1 , eq2 ⟫ = ⟪ begin | |
267 suc (count0 t ) ≡⟨ cong suc eq1 ⟩ | |
268 suc (proj1 (inputnn1-i0 0 t) + count0 (proj2 (inputnn1-i0 0 t))) ≡⟨ cong₂ _+_ (sym (nn31 t 0)) (cong count0 (sym (nn30 t 0))) ⟩ | |
392 | 269 proj1 (inputnn1-i0 1 t) + count0 (proj2 (inputnn1-i0 1 t)) ∎ |
388 | 270 , trans eq2 (cong count1 (sym (nn30 t 0))) ⟫ where |
271 open ≡-Reasoning | |
272 nn20 : proj2 (inputnn1-i0 1 t) ≡ proj2 (inputnn1-i0 0 t) | |
273 nn20 = nn30 t 0 | |
274 nn17 (i1 ∷ x₁) = ⟪ refl , refl ⟫ | |
275 nn19 : (n : ℕ) → (y : List In2 ) → inputnn1-i1 n y ≡ true → (count0 y ≡ 0) ∧ (count1 y ≡ n) | |
276 nn19 zero [] eq = ⟪ refl , refl ⟫ | |
277 nn19 zero (i0 ∷ y) () | |
278 nn19 zero (i1 ∷ y) () | |
392 | 279 nn19 (suc i) (i1 ∷ y) eq with nn19 i y eq |
388 | 280 ... | t = ⟪ proj1 t , cong suc (proj2 t ) ⟫ |
281 nn18 : count0 x ≡ count1 x | |
282 nn18 = begin | |
392 | 283 count0 x ≡⟨ proj1 (nn17 x) ⟩ |
284 proj1 (inputnn1-i0 0 x) + count0 (proj2 (inputnn1-i0 0 x)) ≡⟨ cong (λ k → proj1 (inputnn1-i0 0 x) + k) | |
285 (proj1 (nn19 (proj1 (inputnn1-i0 0 x)) (proj2 (inputnn1-i0 0 x)) eq)) ⟩ | |
286 proj1 (inputnn1-i0 0 x) + 0 ≡⟨ +-comm _ 0 ⟩ | |
287 0 + proj1 (inputnn1-i0 0 x) ≡⟨ cong (λ k → 0 + k) (sym (proj2 (nn19 _ _ eq))) ⟩ | |
288 0 + count1 (proj2 (inputnn1-i0 0 x)) ≡⟨ sym (proj2 (nn17 x)) ⟩ | |
388 | 289 count1 x ∎ where open ≡-Reasoning |
392 | 290 distr0 : (x y : List In2 ) → count0 (x ++ y ) ≡ count0 x + count0 y |
291 distr0 [] y = refl | |
292 distr0 (i0 ∷ x) y = cong suc (distr0 x y) | |
293 distr0 (i1 ∷ x) y = distr0 x y | |
294 distr1 : (x y : List In2 ) → count1 (x ++ y ) ≡ count1 x + count1 y | |
295 distr1 [] y = refl | |
296 distr1 (i1 ∷ x) y = cong suc (distr1 x y) | |
297 distr1 (i0 ∷ x) y = distr1 x y | |
298 -- | |
299 -- i0 .. i0 ∷ i1 .. i1 sequece does not contains i1 → i0 transition | |
300 -- | |
385 | 301 record i1i0 (z : List In2) : Set where |
302 field | |
303 a b : List In2 | |
304 i10 : z ≡ a ++ (i1 ∷ i0 ∷ b ) | |
388 | 305 nn12 : (x : List In2 ) → inputnn1 x ≡ true → ¬ i1i0 x |
391 | 306 nn12 x eq = nn17 x 0 eq where |
392 | 307 nn17 : (x : List In2 ) → (i : ℕ) |
391 | 308 → inputnn1-i1 (proj1 (inputnn1-i0 i x)) (proj2 (inputnn1-i0 i x)) ≡ true → ¬ i1i0 x |
390 | 309 nn17 [] i eq li with i1i0.a li | i1i0.i10 li |
388 | 310 ... | [] | () |
311 ... | x ∷ s | () | |
392 | 312 nn17 (i0 ∷ x₁) i eq li = nn17 x₁ (suc i) eq record { a = nn18 (i1i0.a li) (i1i0.i10 li) ; b = i1i0.b li |
391 | 313 ; i10 = nn19 (i1i0.a li) (i1i0.i10 li) } where |
392 | 314 -- first half |
391 | 315 nn18 : (a : List In2 ) → i0 ∷ x₁ ≡ a ++ ( i1 ∷ i0 ∷ i1i0.b li) → List In2 |
316 nn18 (i0 ∷ t) eq = t | |
317 nn19 : (a : List In2 ) → (eq : i0 ∷ x₁ ≡ a ++ ( i1 ∷ i0 ∷ i1i0.b li) ) | |
318 → x₁ ≡ nn18 a eq ++ i1 ∷ i0 ∷ i1i0.b li | |
319 nn19 (i0 ∷ a) eq = cons-inject eq | |
320 nn17 (i1 ∷ x₁) i eq li = nn20 (i1 ∷ x₁) i eq li where | |
392 | 321 -- second half |
391 | 322 nn20 : (x : List In2) → (i : ℕ) → inputnn1-i1 i x ≡ true → i1i0 x → ⊥ |
392 | 323 nn20 x i eq li = nn21 (i1i0.a li) x i eq (i1i0.i10 li) where |
324 nn21 : (a x : List In2) → (i : ℕ) → inputnn1-i1 i x ≡ true → x ≡ a ++ i1 ∷ i0 ∷ i1i0.b li → ⊥ | |
325 nn21 [] [] zero eq1 () | |
326 nn21 (i0 ∷ a) [] zero eq1 () | |
327 nn21 (i1 ∷ a) [] zero eq1 () | |
328 nn21 a (i0 ∷ x₁) zero () eq0 | |
329 nn21 [] (i0 ∷ x₁) (suc i) () eq0 | |
330 nn21 (x ∷ a) (i0 ∷ x₁) (suc i) () eq0 | |
331 nn21 [] (i1 ∷ i0 ∷ x₁) (suc zero) () eq0 | |
332 nn21 [] (i1 ∷ i0 ∷ x₁) (suc (suc i)) () eq0 | |
333 nn21 (i1 ∷ a) (i1 ∷ x₁) (suc i) eq1 eq0 = nn21 a x₁ i eq1 (cons-inject eq0) | |
387 | 334 nn11 : (x y z : List In2 ) → ¬ y ≡ [] → inputnn1 (x ++ y ++ z) ≡ true → ¬ ( inputnn1 (x ++ y ++ y ++ z) ≡ true ) |
392 | 335 nn11 x y z ny xyz xyyz = ⊥-elim ( nn12 (x ++ y ++ y ++ z ) xyyz record { a = x ++ i1i0.a (bb23 bb22 ) |
388 | 336 ; b = i1i0.b (bb23 bb22) ++ z ; i10 = bb24 } ) where |
392 | 337 -- |
338 -- we need simple calcuraion to obtain count0 y ≡ count1 y | |
339 -- | |
385 | 340 nn21 : count0 x + count0 y + count0 y + count0 z ≡ count1 x + count1 y + count1 y + count1 z |
341 nn21 = begin | |
342 (count0 x + count0 y + count0 y) + count0 z ≡⟨ solve +-0-monoid ⟩ | |
392 | 343 count0 x + (count0 y + (count0 y + count0 z)) ≡⟨ sym (cong (λ k → count0 x + (count0 y + k)) (distr0 y _ )) ⟩ |
344 count0 x + (count0 y + count0 (y ++ z)) ≡⟨ sym (cong (λ k → count0 x + k) (distr0 y _ )) ⟩ | |
345 count0 x + (count0 (y ++ y ++ z)) ≡⟨ sym (distr0 x _ ) ⟩ | |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
346 count0 (x ++ y ++ y ++ z) ≡⟨ nn15 (x ++ y ++ y ++ z) xyyz ⟩ |
392 | 347 count1 (x ++ y ++ y ++ z) ≡⟨ distr1 x _ ⟩ |
348 count1 x + (count1 (y ++ y ++ z)) ≡⟨ cong (λ k → count1 x + k) (distr1 y _ ) ⟩ | |
349 count1 x + (count1 y + count1 (y ++ z)) ≡⟨ (cong (λ k → count1 x + (count1 y + k)) (distr1 y _ )) ⟩ | |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
350 count1 x + (count1 y + (count1 y + count1 z)) ≡⟨ solve +-0-monoid ⟩ |
385 | 351 count1 x + count1 y + count1 y + count1 z ∎ where open ≡-Reasoning |
352 nn20 : count0 x + count0 y + count0 z ≡ count1 x + count1 y + count1 z | |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
353 nn20 = begin |
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
354 count0 x + count0 y + count0 z ≡⟨ solve +-0-monoid ⟩ |
392 | 355 count0 x + (count0 y + count0 z) ≡⟨ cong (λ k → count0 x + k) (sym (distr0 y z)) ⟩ |
356 count0 x + (count0 (y ++ z)) ≡⟨ sym (distr0 x _) ⟩ | |
357 count0 (x ++ (y ++ z)) ≡⟨ nn15 (x ++ y ++ z) xyz ⟩ | |
358 count1 (x ++ (y ++ z)) ≡⟨ distr1 x _ ⟩ | |
359 count1 x + count1 (y ++ z) ≡⟨ cong (λ k → count1 x + k) (distr1 y z) ⟩ | |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
360 count1 x + (count1 y + count1 z) ≡⟨ solve +-0-monoid ⟩ |
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
361 count1 x + count1 y + count1 z ∎ where open ≡-Reasoning |
387 | 362 -- this takes very long time to check and needs 10GB |
385 | 363 bb22 : count0 y ≡ count1 y |
393 | 364 bb22 = ? |
392 | 365 -- |
366 -- y contains i0 and i1 , so we have i1 → i0 transition in y ++ y | |
367 -- | |
385 | 368 bb23 : count0 y ≡ count1 y → i1i0 (y ++ y) |
387 | 369 bb23 eq = bb25 y y bb26 (subst (λ k → 0 < k ) (sym eq) bb26) where |
370 bb26 : 0 < count1 y | |
371 bb26 with <-cmp 0 (count1 y) | |
372 ... | tri< a ¬b ¬c = a | |
373 ... | tri≈ ¬a b ¬c = ⊥-elim (nat-≡< (sym bb27 ) (0<ly y ny) ) where | |
392 | 374 0<ly : (y : List In2) → ¬ y ≡ [] → 0 < length y |
387 | 375 0<ly [] ne = ⊥-elim ( ne refl ) |
376 0<ly (x ∷ y) ne = s≤s z≤n | |
377 bb27 : length y ≡ 0 | |
378 bb27 = begin | |
379 length y ≡⟨ sym (c0+1=n y) ⟩ | |
380 count0 y + count1 y ≡⟨ cong (λ k → k + count1 y ) eq ⟩ | |
381 count1 y + count1 y ≡⟨ cong₂ _+_ (sym b) (sym b) ⟩ | |
382 0 ∎ where open ≡-Reasoning | |
383 bb25 : (x y : List In2 ) → 0 < count1 x → 0 < count0 y → i1i0 (x ++ y) | |
384 bb25 (i0 ∷ x₁) y 0<x 0<y with bb25 x₁ y 0<x 0<y | |
385 ... | t = record { a = i0 ∷ i1i0.a t ; b = i1i0.b t ; i10 = cong (i0 ∷_) (i1i0.i10 t) } | |
386 bb25 (i1 ∷ []) y 0<x 0<y = bb27 y 0<y where | |
387 bb27 : (y : List In2 ) → 0 < count0 y → i1i0 (i1 ∷ y ) | |
392 | 388 bb27 (i0 ∷ y) 0<y = record { a = [] ; b = y ; i10 = refl } |
387 | 389 bb27 (i1 ∷ y) 0<y with bb27 y 0<y |
390 ... | t = record { a = i1 ∷ i1i0.a t ; b = i1i0.b t ; i10 = cong (i1 ∷_) (i1i0.i10 t) } | |
391 bb25 (i1 ∷ i0 ∷ x₁) y 0<x 0<y = record { a = [] ; b = x₁ ++ y ; i10 = refl } | |
392 bb25 (i1 ∷ i1 ∷ x₁) y (s≤s z≤n) 0<y with bb25 (i1 ∷ x₁) y (s≤s z≤n) 0<y | |
393 ... | t = record { a = i1 ∷ i1i0.a t ; b = i1i0.b t ; i10 = cong (i1 ∷_) (i1i0.i10 t) } | |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
394 bb24 : x ++ y ++ y ++ z ≡ (x ++ i1i0.a (bb23 bb22)) ++ i1 ∷ i0 ∷ i1i0.b (bb23 bb22) ++ z |
385 | 395 bb24 = begin |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
396 x ++ y ++ y ++ z ≡⟨ solve (++-monoid In2) ⟩ |
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
397 x ++ (y ++ y) ++ z ≡⟨ cong (λ k → x ++ k ++ z) (i1i0.i10 (bb23 bb22)) ⟩ |
387 | 398 x ++ (i1i0.a (bb23 bb22) ++ i1 ∷ i0 ∷ i1i0.b (bb23 bb22)) ++ z ≡⟨ cong (λ k → x ++ k) -- solver does not work here |
399 (++-assoc (i1i0.a (bb23 bb22)) (i1 ∷ i0 ∷ i1i0.b (bb23 bb22)) z ) ⟩ | |
400 x ++ (i1i0.a (bb23 bb22) ++ (i1 ∷ i0 ∷ i1i0.b (bb23 bb22) ++ z)) ≡⟨ sym (++-assoc x _ _ ) ⟩ | |
386
6ef927ac832c
bb22 takes 10GB and 5 min
Shinji KONO <kono@ie.u-ryukyu.ac.jp>
parents:
385
diff
changeset
|
401 (x ++ i1i0.a (bb23 bb22)) ++ i1 ∷ i0 ∷ i1i0.b (bb23 bb22) ++ z ∎ where open ≡-Reasoning |
385 | 402 |
387 | 403 nn10 : (x y z : List In2 ) → ¬ y ≡ [] → inputnn1 (x ++ y ++ z) ≡ true → inputnn1 (x ++ y ++ y ++ z) ≡ false |
392 | 404 nn10 x y z my eq with inputnn1 (x ++ y ++ y ++ z) | inspect inputnn1 (x ++ y ++ y ++ z) |
387 | 405 ... | true | record { eq = eq1 } = ⊥-elim ( nn11 x y z my eq eq1 ) |
385 | 406 ... | false | _ = refl |
407 | |
304 | 408 |
385 | 409 |
410 | |
411 | |
412 |