File: NNumberSyntax.v

package info (click to toggle)
coq 8.16.1%2Bdfsg-1
  • links: PTS, VCS
  • area: main
  • in suites: bookworm
  • size: 40,596 kB
  • sloc: ml: 219,376; sh: 3,545; python: 3,231; ansic: 2,529; makefile: 767; lisp: 279; javascript: 63; xml: 24; sed: 2
file content (50 lines) | stat: -rw-r--r-- 1,072 bytes parent folder | download | duplicates (2)
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
Require Import NArith.
Check 32%N.
Check (eq_refl : 0x2a%N = 42%N).
Check (fun f : nat -> N => (f 0%nat + 0)%N).
Check (fun x : positive => Npos (xO x)).
Check (fun x : positive => (Npos x + 1)%N).
Check (fun x : positive => Npos x).
Check (fun x : positive => Npos (xI x)).
Check (fun x : positive => (Npos (xO x) + 0)%N).
Check (N.of_nat 0 + 1)%N.
Check (0 + N.of_nat (0 + 0))%N.
Check (N.of_nat 0 = 0%N).
Check 0x00%N : N.
Check 0x01%N : N.
Check 0x02%N : N.
Check 0xff%N : N.
Check 0xFF%N : N.
Check 0x00%xN : N.
Check 0x01%xN : N.
Check 0x02%xN : N.
Check 0xff%xN : N.
Check 0xFF%xN : N.

(* Check hexadecimal printing *)
Open Scope hex_N_scope.
Check 42%N.
Check 0%N.
Check 42%xN.
Check 0%xN.
Check 0x00%N : N.
Check 0x01%N : N.
Check 0x02%N : N.
Check 0xff%N : N.
Check 0xFF%N : N.
Check 0x0%xN : N.
Check 0x00%xN : N.
Check 0x01%xN : N.
Check 0x02%xN : N.
Check 0xff%xN : N.
Check 0xFF%xN : N.
Check 0x0 : N.
Check 0x00 : N.
Check 0x01 : N.
Check 0x02 : N.
Check 0xff : N.
Check 0xFF : N.
Close Scope hex_N_scope.

Require Import Arith.
Check (0 + N.of_nat 11)%N.