@@ -893,7 +893,7 @@ theorem Slice.Pos.offset_str {s : Slice} {pos : s.Pos} :
893893@[simp]
894894theorem Slice.Pos.offset_str_le_offset_endExclusive {s : Slice} {pos : s.Pos} :
895895 pos.str.offset ≤ s.endExclusive.offset := by
896- have := pos.isValidForSlice.le_utf8ByteSize
896+ have := pos.isValidForSlice.le_rawEndPos
897897 have := s.startInclusive_le_endExclusive
898898 simp only [Pos.Raw.le_iff, byteIdx_rawEndPos, utf8ByteSize_eq, offset_str,
899899 Pos.Raw.byteIdx_offsetBy, ValidPos.le_iff] at *
@@ -1012,7 +1012,7 @@ theorem Slice.utf8ByteSize_replaceStartEnd {s : Slice} {newStart newEnd : s.Pos}
10121012theorem Pos.Raw.isValidForSlice_replaceStart {s : Slice} {p : s.Pos} {off : Pos.Raw} :
10131013 off.IsValidForSlice (s.replaceStart p) ↔ (off.offsetBy p.offset).IsValidForSlice s := by
10141014 refine ⟨fun ⟨h₁, h₂⟩ => ⟨?_, ?_⟩, fun ⟨h₁, h₂⟩ => ⟨?_, ?_⟩⟩
1015- · have := p.isValidForSlice.le_utf8ByteSize
1015+ · have := p.isValidForSlice.le_rawEndPos
10161016 simp_all [le_iff]
10171017 omega
10181018 · simpa [Pos.Raw.offsetBy_assoc] using h₂
@@ -1025,7 +1025,7 @@ theorem Pos.Raw.isValidForSlice_replaceEnd {s : Slice} {p : s.Pos} {off : Pos.Ra
10251025 refine ⟨fun ⟨h₁, h₂⟩ => ⟨?_, ?_, ?_⟩, fun ⟨h₁, ⟨h₂, h₃⟩⟩ => ⟨?_, ?_⟩⟩
10261026 · simpa using h₁
10271027 · simp only [Slice.rawEndPos_replaceEnd] at h₁
1028- exact Pos.Raw.le_trans h₁ p.isValidForSlice.le_utf8ByteSize
1028+ exact Pos.Raw.le_trans h₁ p.isValidForSlice.le_rawEndPos
10291029 · simpa using h₂
10301030 · simpa using h₁
10311031 · simpa using h₃
@@ -1297,7 +1297,7 @@ theorem Slice.Pos.ofReplaceStart_startPos {s : Slice} {pos : s.Pos} :
12971297@[simp]
12981298theorem Slice.Pos.ofReplaceStart_endPos {s : Slice} {pos : s.Pos} :
12991299 ofReplaceStart (s.replaceStart pos).endPos = s.endPos := by
1300- have := pos.isValidForSlice.le_utf8ByteSize
1300+ have := pos.isValidForSlice.le_rawEndPos
13011301 simp_all [Pos.ext_iff, String.Pos.Raw.ext_iff, Pos.Raw.le_iff]
13021302
13031303theorem Slice.Pos.ofReplaceStart_inj {s : Slice} {p₀ : s.Pos} {pos pos' : (s.replaceStart p₀).Pos} :
@@ -1388,7 +1388,7 @@ theorem Slice.Pos.lt_next {s : Slice} {pos : s.Pos} {h : pos ≠ s.endPos} :
13881388@[inline, expose]
13891389def Slice.Pos.prevAux {s : Slice} (pos : s.Pos) (h : pos ≠ s.startPos) : String.Pos.Raw :=
13901390 go (pos.offset.byteIdx - 1 ) (by
1391- have := pos.isValidForSlice.le_utf8ByteSize
1391+ have := pos.isValidForSlice.le_rawEndPos
13921392 simp [Pos.Raw.le_iff, Pos.ext_iff] at ⊢ this h
13931393 omega)
13941394where
@@ -1603,7 +1603,7 @@ theorem Slice.Pos.prevAux_lt_self {s : Slice} {p : s.Pos} {h} : p.prevAux h < p.
16031603 omega
16041604
16051605theorem Slice.Pos.prevAux_lt_rawEndPos {s : Slice} {p : s.Pos} {h} : p.prevAux h < s.rawEndPos :=
1606- Pos.Raw.lt_of_lt_of_le prevAux_lt_self p.isValidForSlice.le_utf8ByteSize
1606+ Pos.Raw.lt_of_lt_of_le prevAux_lt_self p.isValidForSlice.le_rawEndPos
16071607
16081608@[simp]
16091609theorem Slice.Pos.prev_ne_endPos {s : Slice} {p : s.Pos} {h} : p.prev h ≠ s.endPos := by
@@ -2256,73 +2256,6 @@ Examples:
22562256@[inline] def splitOn (s : String) (sep : String := " " ) : List String :=
22572257 if sep == "" then [s] else splitOnAux s sep 0 0 0 []
22582258
2259- /--
2260- Adds multiple repetitions of a character to the end of a string.
2261-
2262- Returns `s`, with `n` repetitions of `c` at the end. Internally, the implementation repeatedly calls
2263- `String.push`, so the string is modified in-place if there is a unique reference to it.
2264-
2265- Examples:
2266- * `"indeed".pushn '!' 2 = "indeed!!"`
2267- * `"indeed".pushn '!' 0 = "indeed"`
2268- * `"".pushn ' ' 4 = " "`
2269- -/
2270- @[inline] def pushn (s : String) (c : Char) (n : Nat) : String :=
2271- n.repeat (fun s => s.push c) s
2272-
2273- @[export lean_string_pushn]
2274- def Internal.pushnImpl (s : String) (c : Char) (n : Nat) : String :=
2275- String.pushn s c n
2276-
2277- /--
2278- Checks whether a string is empty.
2279-
2280- Empty strings are equal to `""` and have length and end position `0`.
2281-
2282- Examples:
2283- * `"".isEmpty = true`
2284- * `"empty".isEmpty = false`
2285- * `" ".isEmpty = false`
2286- -/
2287- @[inline] def isEmpty (s : String) : Bool :=
2288- s.rawEndPos == 0
2289-
2290- @[export lean_string_isempty]
2291- def Internal.isEmptyImpl (s : String) : Bool :=
2292- String.isEmpty s
2293-
2294- /--
2295- Appends all the strings in a list of strings, in order.
2296-
2297- Use `String.intercalate` to place a separator string between the strings in a list.
2298-
2299- Examples:
2300- * `String.join ["gr", "ee", "n"] = "green"`
2301- * `String.join ["b", "", "l", "", "ue"] = "blue"`
2302- * `String.join [] = ""`
2303- -/
2304- @[inline] def join (l : List String) : String :=
2305- l.foldl (fun r s => r ++ s) ""
2306-
2307- /--
2308- Appends the strings in a list of strings, placing the separator `s` between each pair.
2309-
2310- Examples:
2311- * `", ".intercalate ["red", "green", "blue"] = "red, green, blue"`
2312- * `" and ".intercalate ["tea", "coffee"] = "tea and coffee"`
2313- * `" | ".intercalate ["M", "", "N"] = "M | | N"`
2314- -/
2315- def intercalate (s : String) : List String → String
2316- | [] => ""
2317- | a :: as => go a s as
2318- where go (acc : String) (s : String) : List String → String
2319- | a :: as => go (acc ++ s ++ a) s as
2320- | [] => acc
2321-
2322- @[export lean_string_intercalate]
2323- def Internal.intercalateImpl (s : String) : List String → String :=
2324- String.intercalate s
2325-
23262259
23272260def offsetOfPosAux (s : String) (pos : Pos.Raw) (i : Pos.Raw) (offset : Nat) : Nat :=
23282261 if i >= pos then offset
@@ -2592,7 +2525,7 @@ theorem length_singleton {c : Char} : (String.singleton c).length = 1 := by
25922525 simp [← length_data]
25932526
25942527@[simp] theorem length_pushn (c : Char) (n : Nat) : (pushn s c n).length = s.length + n := by
2595- unfold pushn ; induction n <;> simp [Nat.repeat, Nat.add_assoc, *]
2528+ rw [pushn_eq_repeat_push] ; induction n <;> simp [Nat.repeat, Nat.add_assoc, *]
25962529
25972530@[simp] theorem length_append (s t : String) : (s ++ t).length = s.length + t.length := by
25982531 simp [← length_data]
0 commit comments