utf8EncodeChar #
utf8EncodeChar low-level API #
Returns the sequence of bytes in a character's UTF-8 encoding.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[csimp]
@[simp]
utf8EncodeChar BitVec API #
Size one #
Size two #
Size three #
Size four #
parseFirstByte definition #
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
parseFirstByte low-level API #
parseFirstByte BitVec API #
Size one #
Size two #
Size three #
Size four #
isInvalidContinuationByte definition & API #
@[inline]
Equations
- ByteArray.utf8DecodeChar?.isInvalidContinuationByte b = (b &&& 192 != 128)
Instances For
parseFirstByte, isInvalidContinuationByte and utf8EncodeChar #
@[inline]
Instances For
@[inline]
def
ByteArray.utf8DecodeChar?.verify₁
{w : UInt8}
(_w : UInt8)
(_h : parseFirstByte w = FirstByte.done)
:
Equations
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Decodes and returns the Char whose UTF-8 encoding begins at i in bytes.
Returns none if i is not the start of a valid UTF-8 encoding of a character.
Equations
- One or more equations did not get rendered due to their size.
Instances For
@[inline]
Equations
- One or more equations did not get rendered due to their size.
Instances For
utf8DecodeChar? low-level API #
Main theorems #
utf8DecodeChar?_utf8EncodeChar_append and toByteArray_of_utf8DecodeChar?_eq_some are the two main results that together
imply that UTF-8 encoding and decoding are inverse.
theorem
String.toByteArray_utf8EncodeChar_of_utf8DecodeChar?_eq_some
{b : ByteArray}
{c : Char}
(h : b.utf8DecodeChar? 0 = some c)
:
Corollaries #
theorem
ByteArray.eq_of_utf8DecodeChar?_eq_some
{b : ByteArray}
{c : Char}
(h : b.utf8DecodeChar? 0 = some c)
:
theorem
ByteArray.exists_of_utf8DecodeChar?_eq_some
{b : ByteArray}
{c : Char}
(h : b.utf8DecodeChar? 0 = some c)
:
theorem
ByteArray.lt_size_of_isSome_utf8DecodeChar?
{b : ByteArray}
{i : Nat}
(h : (b.utf8DecodeChar? i).isSome = true)
:
theorem
ByteArray.lt_size_of_validateUTF8At
{b : ByteArray}
{i : Nat}
:
b.validateUTF8At i = true → i < b.size
theorem
ByteArray.utf8DecodeChar?_append_eq_some
{b : ByteArray}
{i : Nat}
{c : Char}
(h : b.utf8DecodeChar? i = some c)
(b' : ByteArray)
:
theorem
ByteArray.isSome_utf8DecodeChar?_append
{b : ByteArray}
{i : Nat}
(h : (b.utf8DecodeChar? i).isSome = true)
(b' : ByteArray)
:
@[inline]
def
ByteArray.utf8DecodeChar
(bytes : ByteArray)
(i : Nat)
(h : (bytes.utf8DecodeChar? i).isSome = true)
:
Decodes and returns the Char whose UTF-8 encoding begins at i in bytes.
This function requires a proof that there is, in fact, a valid Char at i. utf8DecodeChar? is
an alternative function that returns Option Char instead of requiring a proof ahead of time.
Equations
- bytes.utf8DecodeChar i h = (bytes.utf8DecodeChar? i).get h
Instances For
theorem
ByteArray.add_utf8Size_utf8DecodeChar_le_size
{b : ByteArray}
{i : Nat}
{h : (b.utf8DecodeChar? i).isSome = true}
:
theorem
ByteArray.utf8DecodeChar_eq_utf8DecodeChar_extract
{b : ByteArray}
{i : Nat}
{h : (b.utf8DecodeChar? i).isSome = true}
:
theorem
ByteArray.utf8DecodeChar_extract_congr
{bytes : ByteArray}
(i j j' : Nat)
{h : ((bytes.extract i j).utf8DecodeChar? 0).isSome = true}
{h' : ((bytes.extract i j').utf8DecodeChar? 0).isSome = true}
:
theorem
ByteArray.utf8EncodeChar_utf8DecodeChar
{b : ByteArray}
{i : Nat}
{h : (b.utf8DecodeChar? i).isSome = true}
:
(String.utf8EncodeChar (b.utf8DecodeChar i h)).toByteArray = b.extract i (i + (b.utf8DecodeChar i h).utf8Size)
@[simp]
@[simp]
@[simp]
@[simp]
theorem
List.utf8DecodeChar_utf8Encode_cons
{l : List Char}
{c : Char}
{h : ((c :: l).utf8Encode.utf8DecodeChar? 0).isSome = true}
:
@[simp]
theorem
UInt8.isUTF8FirstByte_getElem_utf8EncodeChar
{c : Char}
{i : Nat}
{hi : i < (String.utf8EncodeChar c).length}
:
theorem
ByteArray.isUTF8FirstByte_of_isSome_utf8DecodeChar?
{b : ByteArray}
{i : Nat}
(h : (b.utf8DecodeChar? i).isSome = true)
:
theorem
ByteArray.isUTF8FirstByte_of_validateUTF8At
{b : ByteArray}
{i : Nat}
(h : b.validateUTF8At i = true)
:
theorem
ByteArray.utf8Size_utf8DecodeChar
{b : ByteArray}
{i : Nat}
{h : (b.utf8DecodeChar? i).isSome = true}
: