summaryrefslogtreecommitdiff
path: root/theories/Strings/HexString.v
blob: 9ea93c909e32ae009e68c141f34d67b404dd72c8 (plain)
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
(************************************************************************)
(*  v      *   The Coq Proof Assistant  /  The Coq Development Team     *)
(* <O___,, *   INRIA - CNRS - LIX - LRI - PPS - Copyright 1999-2017     *)
(*   \VV/  **************************************************************)
(*    //   *      This file is distributed under the terms of the       *)
(*         *       GNU Lesser General Public License Version 2.1        *)
(************************************************************************)

Require Import Ascii String.
Require Import BinNums.
Import BinNatDef.
Import BinIntDef.
Import BinPosDef.

Local Open Scope positive_scope.
Local Open Scope string_scope.

Local Notation "a || b"
  := (if a then true else if b then true else false).
Definition ascii_to_digit (ch : ascii) : option N
  := (if ascii_dec ch "0" then Some 0
      else if ascii_dec ch "1" then Some 1
      else if ascii_dec ch "2" then Some 2
      else if ascii_dec ch "3" then Some 3
      else if ascii_dec ch "4" then Some 4
      else if ascii_dec ch "5" then Some 5
      else if ascii_dec ch "6" then Some 6
      else if ascii_dec ch "7" then Some 7
      else if ascii_dec ch "8" then Some 8
      else if ascii_dec ch "9" then Some 9
      else if ascii_dec ch "a" || ascii_dec ch "A" then Some 10
      else if ascii_dec ch "b" || ascii_dec ch "B" then Some 11
      else if ascii_dec ch "c" || ascii_dec ch "C" then Some 12
      else if ascii_dec ch "d" || ascii_dec ch "D" then Some 13
      else if ascii_dec ch "e" || ascii_dec ch "E" then Some 14
      else if ascii_dec ch "f" || ascii_dec ch "F" then Some 15
      else None)%N.

Fixpoint pos_hex_app (p q:positive) : positive :=
  match q with
  | 1 => p~0~0~0~1
  | 2 => p~0~0~1~0
  | 3 => p~0~0~1~1
  | 4 => p~0~1~0~0
  | 5 => p~0~1~0~1
  | 6 => p~0~1~1~0
  | 7 => p~0~1~1~1
  | 8 => p~1~0~0~0
  | 9 => p~1~0~0~1
  | 10 => p~1~0~1~0
  | 11 => p~1~0~1~1
  | 12 => p~1~1~0~0
  | 13 => p~1~1~0~1
  | 14 => p~1~1~1~0
  | 15 => p~1~1~1~1
  | q~0~0~0~0 => (pos_hex_app p q)~0~0~0~0
  | q~0~0~0~1 => (pos_hex_app p q)~0~0~0~1
  | q~0~0~1~0 => (pos_hex_app p q)~0~0~1~0
  | q~0~0~1~1 => (pos_hex_app p q)~0~0~1~1
  | q~0~1~0~0 => (pos_hex_app p q)~0~1~0~0
  | q~0~1~0~1 => (pos_hex_app p q)~0~1~0~1
  | q~0~1~1~0 => (pos_hex_app p q)~0~1~1~0
  | q~0~1~1~1 => (pos_hex_app p q)~0~1~1~1
  | q~1~0~0~0 => (pos_hex_app p q)~1~0~0~0
  | q~1~0~0~1 => (pos_hex_app p q)~1~0~0~1
  | q~1~0~1~0 => (pos_hex_app p q)~1~0~1~0
  | q~1~0~1~1 => (pos_hex_app p q)~1~0~1~1
  | q~1~1~0~0 => (pos_hex_app p q)~1~1~0~0
  | q~1~1~0~1 => (pos_hex_app p q)~1~1~0~1
  | q~1~1~1~0 => (pos_hex_app p q)~1~1~1~0
  | q~1~1~1~1 => (pos_hex_app p q)~1~1~1~1
  end.

Module Raw.
  Fixpoint of_pos (p : positive) (rest : string) : string
    := match p with
       | 1 => String "1" rest
       | 2 => String "2" rest
       | 3 => String "3" rest
       | 4 => String "4" rest
       | 5 => String "5" rest
       | 6 => String "6" rest
       | 7 => String "7" rest
       | 8 => String "8" rest
       | 9 => String "9" rest
       | 10 => String "a" rest
       | 11 => String "b" rest
       | 12 => String "c" rest
       | 13 => String "d" rest
       | 14 => String "e" rest
       | 15 => String "f" rest
       | p'~0~0~0~0 => of_pos p' (String "0" rest)
       | p'~0~0~0~1 => of_pos p' (String "1" rest)
       | p'~0~0~1~0 => of_pos p' (String "2" rest)
       | p'~0~0~1~1 => of_pos p' (String "3" rest)
       | p'~0~1~0~0 => of_pos p' (String "4" rest)
       | p'~0~1~0~1 => of_pos p' (String "5" rest)
       | p'~0~1~1~0 => of_pos p' (String "6" rest)
       | p'~0~1~1~1 => of_pos p' (String "7" rest)
       | p'~1~0~0~0 => of_pos p' (String "8" rest)
       | p'~1~0~0~1 => of_pos p' (String "9" rest)
       | p'~1~0~1~0 => of_pos p' (String "a" rest)
       | p'~1~0~1~1 => of_pos p' (String "b" rest)
       | p'~1~1~0~0 => of_pos p' (String "c" rest)
       | p'~1~1~0~1 => of_pos p' (String "d" rest)
       | p'~1~1~1~0 => of_pos p' (String "e" rest)
       | p'~1~1~1~1 => of_pos p' (String "f" rest)
       end.

  Fixpoint to_N (s : string) (rest : N)
    : N
    := match s with
       | "" => rest
       | String ch s'
         => to_N
              s'
              match ascii_to_digit ch with
              | Some v => N.add v (N.mul 16 rest)
              | None => N0
              end
       end.

  Fixpoint to_N_of_pos (p : positive) (rest : string) (base : N)
    : to_N (of_pos p rest) base
      = to_N rest match base with
                  | N0 => N.pos p
                  | Npos v => Npos (pos_hex_app v p)
                  end.
  Proof.
    do 4 try destruct p as [p|p|]; destruct base; try reflexivity;
      cbn; rewrite to_N_of_pos; reflexivity.
  Qed.
End Raw.

Definition of_pos (p : positive) : string
  := String "0" (String "x" (Raw.of_pos p "")).
Definition of_N (n : N) : string
  := match n with
     | N0 => "0x0"
     | Npos p => of_pos p
     end.
Definition of_Z (z : Z) : string
  := match z with
     | Zneg p => String "-" (of_pos p)
     | Z0 => "0x0"
     | Zpos p => of_pos p
     end.
Definition of_nat (n : nat) : string
  := of_N (N.of_nat n).

Definition to_N (s : string) : N
  := match s with
     | String s0 (String so s)
       => if ascii_dec s0 "0"
          then if ascii_dec so "x"
               then Raw.to_N s N0
               else N0
          else N0
     | _ => N0
     end.
Definition to_pos (s : string) : positive
  := match to_N s with
     | N0 => 1
     | Npos p => p
     end.
Definition to_Z (s : string) : Z
  := let '(is_neg, n) := match s with
                         | String s0 s'
                           => if ascii_dec s0 "-"
                              then (true, to_N s')
                              else (false, to_N s)
                         | EmptyString => (false, to_N s)
                         end in
     match n with
     | N0 => Z0
     | Npos p => if is_neg then Zneg p else Zpos p
     end.
Definition to_nat (s : string) : nat
  := N.to_nat (to_N s).

Lemma to_N_of_N (n : N)
  : to_N (of_N n)
    = n.
Proof.
  destruct n; [ reflexivity | apply Raw.to_N_of_pos ].
Qed.

Lemma to_Z_of_Z (z : Z)
  : to_Z (of_Z z)
    = z.
Proof.
  cbv [of_Z to_Z]; destruct z as [|z|z]; cbn;
    try reflexivity;
    rewrite Raw.to_N_of_pos; cbn; reflexivity.
Qed.

Lemma to_nat_of_nat (n : nat)
  : to_nat (of_nat n)
    = n.
Proof.
  cbv [to_nat of_nat];
    rewrite to_N_of_N, Nnat.Nat2N.id; reflexivity.
Qed.

Lemma to_pos_of_pos (p : positive)
  : to_pos (of_pos p)
    = p.
Proof.
  cbv [of_pos to_pos to_N]; cbn;
    rewrite Raw.to_N_of_pos; cbn; reflexivity.
Qed.

Example of_pos_1 : of_pos 1 = "0x1" := eq_refl.
Example of_pos_2 : of_pos 2 = "0x2" := eq_refl.
Example of_pos_3 : of_pos 3 = "0x3" := eq_refl.
Example of_pos_7 : of_pos 7 = "0x7" := eq_refl.
Example of_pos_8 : of_pos 8 = "0x8" := eq_refl.
Example of_pos_9 : of_pos 9 = "0x9" := eq_refl.
Example of_pos_10 : of_pos 10 = "0xa" := eq_refl.
Example of_pos_11 : of_pos 11 = "0xb" := eq_refl.
Example of_pos_12 : of_pos 12 = "0xc" := eq_refl.
Example of_pos_13 : of_pos 13 = "0xd" := eq_refl.
Example of_pos_14 : of_pos 14 = "0xe" := eq_refl.
Example of_pos_15 : of_pos 15 = "0xf" := eq_refl.
Example of_pos_16 : of_pos 16 = "0x10" := eq_refl.
Example of_N_0 : of_N 0 = "0x0" := eq_refl.
Example of_Z_0 : of_Z 0 = "0x0" := eq_refl.
Example of_Z_m1 : of_Z (-1) = "-0x1" := eq_refl.
Example of_nat_0 : of_nat 0 = "0x0" := eq_refl.