diff options
| author | Pierre Letouzey | 2018-03-08 18:00:47 +0100 |
|---|---|---|
| committer | Jason Gross | 2018-07-16 09:50:42 -0400 |
| commit | a697b20f5013abb834cd9dca3fdef2bee53221ad (patch) | |
| tree | abb0d353247d2bdc85d9d3dc919962d993ab27ae /theories | |
| parent | e27d2fd32e40937e48889772361f1cf53dd9c48e (diff) | |
Ascii.eqb and String.eqb
Diffstat (limited to 'theories')
| -rw-r--r-- | theories/Bool/Bool.v | 7 | ||||
| -rw-r--r-- | theories/Numbers/DecimalString.v | 20 | ||||
| -rw-r--r-- | theories/Strings/Ascii.v | 19 | ||||
| -rw-r--r-- | theories/Strings/String.v | 19 |
4 files changed, 55 insertions, 10 deletions
diff --git a/theories/Bool/Bool.v b/theories/Bool/Bool.v index edf78ed52d..66a82008d8 100644 --- a/theories/Bool/Bool.v +++ b/theories/Bool/Bool.v @@ -814,3 +814,10 @@ Defined. (** Reciprocally, from a decidability, we could state a [reflect] as soon as we have a [bool_of_sumbool]. *) + +(** For instance, we could state the correctness of [Bool.eqb] via [reflect]: *) + +Lemma eqb_spec (b b' : bool) : reflect (b = b') (eqb b b'). +Proof. + destruct b, b'; now constructor. +Qed. diff --git a/theories/Numbers/DecimalString.v b/theories/Numbers/DecimalString.v index 1a3220f63a..591024baec 100644 --- a/theories/Numbers/DecimalString.v +++ b/theories/Numbers/DecimalString.v @@ -94,7 +94,7 @@ Definition int_of_string s := match s with | EmptyString => Some (Pos Nil) | String a s' => - if ascii_dec a "-" then option_map Neg (uint_of_string s') + if Ascii.eqb a "-" then option_map Neg (uint_of_string s') else option_map Pos (uint_of_string s) end. @@ -131,8 +131,8 @@ Proof. - unfold int_of_string. destruct (string_of_uint d) eqn:Hd. + now destruct d. - + destruct ascii_dec; subst. - * now destruct d. + + case Ascii.eqb_spec. + * intros ->. now destruct d. * rewrite <- Hd, usu; auto. - rewrite usu; auto. Qed. @@ -141,8 +141,8 @@ Lemma sis s d : int_of_string s = Some d -> string_of_int d = s. Proof. destruct s; [intros [= <-]| ]; simpl; trivial. - destruct ascii_dec; subst; simpl. - - destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-]. + case Ascii.eqb_spec. + - intros ->. destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-]. simpl; f_equal. now apply sus. - destruct d; [ | now destruct uint_of_char]. simpl string_of_int. @@ -178,7 +178,7 @@ Definition int_of_string s := match s with | EmptyString => None | String a s' => - if ascii_dec a "-" then option_map Neg (uint_of_string s') + if Ascii.eqb a "-" then option_map Neg (uint_of_string s') else option_map Pos (uint_of_string s) end. @@ -228,8 +228,8 @@ Proof. unfold int_of_string. destruct (string_of_uint d) eqn:Hd. + now destruct d. - + destruct ascii_dec; subst. - * now destruct d. + + case Ascii.eqb_spec. + * intros ->. now destruct d. * rewrite <- Hd, usu; auto. now intros ->. - intros _ H. rewrite usu; auto. now intros ->. @@ -253,8 +253,8 @@ Lemma sis s d : int_of_string s = Some d -> string_of_int d = s. Proof. destruct s; [intros [=]| ]; simpl. - destruct ascii_dec; subst; simpl. - - destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-]. + case Ascii.eqb_spec. + - intros ->. destruct (uint_of_string s) eqn:Hs; simpl; intros [= <-]. simpl; f_equal. now apply sus. - destruct d; [ | now destruct uint_of_char]. simpl string_of_int. diff --git a/theories/Strings/Ascii.v b/theories/Strings/Ascii.v index 5154b75b3f..51da2cc61b 100644 --- a/theories/Strings/Ascii.v +++ b/theories/Strings/Ascii.v @@ -40,6 +40,25 @@ Proof. decide equality; apply bool_dec. Defined. +Local Open Scope lazy_bool_scope. + +Definition eqb (a b : ascii) : bool := + match a, b with + | Ascii a0 a1 a2 a3 a4 a5 a6 a7, + Ascii b0 b1 b2 b3 b4 b5 b6 b7 => + Bool.eqb a0 b0 &&& Bool.eqb a1 b1 &&& Bool.eqb a2 b2 &&& Bool.eqb a3 b3 + &&& Bool.eqb a4 b4 &&& Bool.eqb a5 b5 &&& Bool.eqb a6 b6 &&& Bool.eqb a7 b7 + end. + +Infix "=?" := eqb : char_scope. + +Lemma eqb_spec (a b : ascii) : reflect (a = b) (a =? b)%char. +Proof. + destruct a, b; simpl. + do 8 (case Bool.eqb_spec; [ intros -> | constructor; now intros [= ] ]). + now constructor. +Qed. + (** * Conversion between natural numbers modulo 256 and ascii characters *) (** Auxiliary function that turns a positive into an ascii by diff --git a/theories/Strings/String.v b/theories/Strings/String.v index 2be6618ad6..78268eeed8 100644 --- a/theories/Strings/String.v +++ b/theories/Strings/String.v @@ -14,6 +14,7 @@ Require Import Arith. Require Import Ascii. +Require Import Bool. Declare ML Module "string_syntax_plugin". (** *** Definition of strings *) @@ -35,6 +36,24 @@ Proof. decide equality; apply ascii_dec. Defined. +Local Open Scope lazy_bool_scope. + +Fixpoint eqb s1 s2 : bool := + match s1, s2 with + | EmptyString, EmptyString => true + | String c1 s1', String c2 s2' => Ascii.eqb c1 c2 &&& eqb s1' s2' + | _,_ => false + end. + +Infix "=?" := eqb : string_scope. + +Lemma eqb_spec s1 s2 : Bool.reflect (s1 = s2) (s1 =? s2)%string. +Proof. + revert s2. induction s1; destruct s2; try (constructor; easy); simpl. + case Ascii.eqb_spec; simpl; [intros -> | constructor; now intros [= ]]. + case IHs1; [intros ->; now constructor | constructor; now intros [= ]]. +Qed. + (** *** Concatenation of strings *) Reserved Notation "x ++ y" (right associativity, at level 60). |
