@@ -3,8 +3,10 @@ Copyright (c) 2023 Bulhwi Cha. All rights reserved.
3
3
Released under Apache 2.0 license as described in the file LICENSE.
4
4
Authors: Bulhwi Cha, Mario Carneiro
5
5
-/
6
+ import Batteries.Logic
6
7
import Batteries.Data.Char
7
- import Batteries.Data.List.Lemmas
8
+ import Batteries.Data.Nat.Lemmas
9
+ import Batteries.Data.List.SplitOnList
8
10
import Batteries.Data.String.Basic
9
11
import Batteries.Tactic.Lint.Misc
10
12
import Batteries.Tactic.SeqFocus
@@ -440,7 +442,83 @@ theorem splitAux_of_valid (p l m r acc) :
440
442
theorem split_of_valid (s p) : split s p = (List.splitOnP p s.1 ).map mk := by
441
443
simpa [split] using splitAux_of_valid p [] [] s.1 []
442
444
443
- -- TODO: splitOn
445
+ theorem splitOnAux_of_valid (sep₁ sep₂ l m r acc) (h : sep₂ ≠ []) :
446
+ splitOnAux ⟨l ++ m ++ sep₁ ++ r⟩ ⟨sep₁ ++ sep₂⟩ ⟨utf8Len l⟩ ⟨utf8Len (l ++ m ++ sep₁)⟩
447
+ ⟨utf8Len sep₁⟩ acc =
448
+ acc.reverse ++ (List.modifyHead (m ++ ·) <| List.splitOnListAux r sep₁ sep₂ h).map mk := by
449
+ obtain ⟨c₂, sep₂, rfl⟩ := sep₂.exists_cons_of_ne_nil h
450
+ rw [splitOnAux, List.splitOnListAux]
451
+ simp only [List.append_assoc, utf8Len_append, atEnd_iff, endPos_eq, Pos.mk_le_mk,
452
+ Nat.add_le_add_iff_left, Nat.add_right_le_self, utf8Len_eq_zero, by simpa only
453
+ [List.append_assoc, utf8Len_append] using (⟨get_of_valid (l++m++sep₁) r, next_of_valid
454
+ (l++m++sep₁) (r.headD default) r, extract_of_valid l (m++sep₁) r⟩ : _∧_∧_), List.reverse_cons,
455
+ get_of_valid, List.headD_cons, beq_iff_eq, next, Pos.addChar_eq, utf8Len_cons,
456
+ Nat.add_left_le_self, Pos.sub_eq, dite_eq_ite]
457
+ split
458
+ · simp
459
+ · next hls =>
460
+ obtain ⟨cᵣ, r, rfl⟩ := r.exists_cons_of_ne_nil hls
461
+ simp only [List.headD_cons, List.head_cons, List.tail_cons, false_or]
462
+ split
463
+ · next hc =>
464
+ subst c₂
465
+ rw [← Nat.add_assoc, Nat.add_assoc, Nat.add_sub_cancel, Nat.add_assoc]
466
+ split
467
+ · subst sep₂
468
+ have ih := by simpa only [List.append_assoc, List.append_nil, List.singleton_append,
469
+ List.nil_append, utf8Len_append, utf8Len_cons, utf8Len_nil, Nat.zero_add, Pos.mk_zero,
470
+ List.reverse_cons, ← Function.id_def, List.modifyHead_id] using (splitOnAux_of_valid []
471
+ (sep₁++[cᵣ]) (l++m++sep₁++[cᵣ]) [] r (⟨m⟩::acc) (by simp))
472
+ simp [by simpa using extract_of_valid l m (sep₁++cᵣ::r), ih]
473
+ · next hsp₂ =>
474
+ simpa using splitOnAux_of_valid (sep₁++[cᵣ]) sep₂ l m r acc hsp₂
475
+ · next hc =>
476
+ have hget := by simpa only [List.append_assoc, utf8Len_append]
477
+ using get_of_valid (l++m) (sep₁++cᵣ::r)
478
+ have ih := by simpa only [List.append_assoc, List.append_nil, List.singleton_append,
479
+ List.nil_append, utf8Len_append, utf8Len_cons, utf8Len_nil, Nat.zero_add, Pos.mk_zero,
480
+ List.head_cons_tail (sep₁++cᵣ::r) (by simp)] using (splitOnAux_of_valid [] (sep₁++c₂::sep₂)
481
+ l (m++[(sep₁++cᵣ::r).head (by simp)]) (sep₁++cᵣ::r).tail acc (by simp))
482
+ rw [Nat.add_sub_assoc, Nat.add_sub_cancel, hget, List.headD_eq_head (sep₁++cᵣ::r) (by simp),
483
+ Nat.add_assoc, ih, List.append_cancel_left_eq, List.modifyHead_modifyHead,
484
+ Function.comp_def]; clear hget ih
485
+ apply Nat.le_add_left
486
+ termination_by (utf8Len sep₁ + utf8Len r, utf8Len sep₂)
487
+ decreasing_by
488
+ all_goals simp_wf
489
+ · rename_i _₀ _₁ _sep₂ _₂ _₃ _₄ _₅ _₆ _₇ _r _₈ _₉ _₁₀ _₁₁ _₁₂ _ _₁₃ _₁₄ _ _hc
490
+ clear _₀ _₁ _₂ _₃ _₄ _₅ _₆ _₇ _₈ _₉ _₁₀ _₁₁ _₁₂ _₁₃ _₁₄
491
+ subst _r _sep₂
492
+ left
493
+ rw [← utf8Len_append]
494
+ conv => rhs; rw [← List.head_cons_tail (sep₁++cᵣ::r) (by simp), utf8Len_cons]
495
+ apply Pos.lt_addChar
496
+ · rename_i _₀ _₁ _sep₂ _₂ _₃ _₄ _₅ _₆ _₇ _r _₈ _₉ _₁₀ _₁₁ _₁₂ _ _₁₃ _₁₄ _₁₅ _ _hc _hsp₂
497
+ clear _₀ _₁ _₂ _₃ _₄ _₅ _₆ _₇ _₈ _₉ _₁₀ _₁₁ _₁₂ _₁₃ _₁₄ _₁₅
498
+ subst _r _sep₂ c₂
499
+ repeat rw [utf8Len_cons]
500
+ rw [Nat.add_right_comm, Nat.add_assoc]
501
+ right
502
+ apply Pos.lt_addChar
503
+ · rename_i _₀ _₁ _sep₂ _₂ _₃ _₄ _₅ _₆ _₇ _r _₈ _₉ _₁₀ _₁₁ _₁₂ _ _₁₃ _₁₄ _₁₅ _ _hc _₁₆ _hsp₂
504
+ clear _₀ _₁ _₂ _₃ _₄ _₅ _₆ _₇ _₈ _₉ _₁₀ _₁₁ _₁₂ _₁₃ _₁₄ _₁₅ _₁₆
505
+ subst _r _sep₂
506
+ left
507
+ rw [utf8Len_cons]
508
+ calc
509
+ utf8Len r < utf8Len r + cᵣ.utf8Size := Pos.lt_addChar ..
510
+ utf8Len r + cᵣ.utf8Size ≤ utf8Len sep₁ + (utf8Len r + cᵣ.utf8Size) := Nat.le_add_left ..
511
+
512
+ theorem splitOn_of_valid (s sep) : splitOn s sep = (List.splitOnList s.1 sep.1 ).map mk := by
513
+ simp only [splitOn, beq_iff_eq]
514
+ split <;> rename_i hsp
515
+ · simp [hsp, List.splitOnList]
516
+ · have aux := by simpa only [List.append_nil, List.nil_append, utf8Len_nil, Pos.mk_zero,
517
+ List.reverse_nil, ← Function.id_def, List.modifyHead_id] using
518
+ splitOnAux_of_valid [] sep.1 [] [] s.1 [] (hsp ∘ ext)
519
+ rw [aux]; clear aux
520
+ apply congrArg
521
+ simp [List.splitOnListAux_eq_splitOnList_append_append]
444
522
445
523
@[simp] theorem toString_toSubstring (s : String) : s.toSubstring.toString = s :=
446
524
extract_zero_endPos _
0 commit comments