src/HOL/Tools/numeral.ML
author desharna
Mon, 10 Oct 2022 19:07:54 +0200
changeset 76256 207b6fcfc47d
parent 74282 c2ee8d993d6a
child 77879 dd222e2af01a
permissions -rw-r--r--
removed unused universal variable from lemma reflp_onI
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
28308
d4396a28fb29 fixed headers
haftmann
parents: 28262
diff changeset
     1
(*  Title:      HOL/Tools/numeral.ML
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
     3
68028
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 62597
diff changeset
     4
Logical and syntactic operations on numerals (see also HOL/Tools/hologic.ML).
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
     5
*)
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
     6
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
     7
signature NUMERAL =
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
     8
sig
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 23575
diff changeset
     9
  val mk_cnumber: ctyp -> int -> cterm
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58399
diff changeset
    10
  val mk_number_syntax: int -> term
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    11
  val dest_num_syntax: term -> int
58399
haftmann
parents: 55236
diff changeset
    12
  val add_code: string -> (int -> int) -> (Code_Printer.literals -> int -> string) -> string -> theory -> theory
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    13
end;
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    14
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    15
structure Numeral: NUMERAL =
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    16
struct
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    17
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    18
(* numeral *)
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    19
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    20
fun dest_num_syntax (Const (\<^const_syntax>\<open>Num.Bit0\<close>, _) $ t) = 2 * dest_num_syntax t
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    21
  | dest_num_syntax (Const (\<^const_syntax>\<open>Num.Bit1\<close>, _) $ t) = 2 * dest_num_syntax t + 1
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    22
  | dest_num_syntax (Const (\<^const_syntax>\<open>Num.One\<close>, _)) = 1;
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    23
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    24
fun mk_num_syntax n =
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    25
  if n > 0 then
73020
b51515722274 tuned signature -- prefer Isabelle/ML structure Integer;
wenzelm
parents: 69593
diff changeset
    26
    (case Integer.quot_rem n 2 of
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    27
      (0, 1) => Syntax.const \<^const_syntax>\<open>One\<close>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    28
    | (n, 0) => Syntax.const \<^const_syntax>\<open>Bit0\<close> $ mk_num_syntax n
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    29
    | (n, 1) => Syntax.const \<^const_syntax>\<open>Bit1\<close> $ mk_num_syntax n)
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    30
  else raise Match
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    31
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    32
fun mk_cbit 0 = \<^cterm>\<open>Num.Bit0\<close>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    33
  | mk_cbit 1 = \<^cterm>\<open>Num.Bit1\<close>
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    34
  | mk_cbit _ = raise CTERM ("mk_cbit", []);
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    35
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    36
fun mk_cnumeral i =
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    37
  let
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    38
    fun mk 1 = \<^cterm>\<open>Num.One\<close>
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    39
      | mk i =
24630
351a308ab58d simplified type int (eliminated IntInf.int, integer);
wenzelm
parents: 23575
diff changeset
    40
      let val (q, r) = Integer.div_mod i 2 in
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    41
        Thm.apply (mk_cbit r) (mk q)
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    42
      end
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    43
  in
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    44
    if i > 0 then mk i else raise CTERM ("mk_cnumeral: negative input", [])
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    45
  end
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    46
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    47
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    48
(* number *)
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    49
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    50
local
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    51
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    52
val cterm_of = Thm.cterm_of \<^context>;
61150
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    53
fun tvar S = (("'a", 0), S);
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    54
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    55
val zero_tvar = tvar \<^sort>\<open>zero\<close>;
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    56
val zero = cterm_of (Const (\<^const_name>\<open>zero_class.zero\<close>, TVar zero_tvar));
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    57
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    58
val one_tvar = tvar \<^sort>\<open>one\<close>;
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    59
val one = cterm_of (Const (\<^const_name>\<open>one_class.one\<close>, TVar one_tvar));
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    60
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    61
val numeral_tvar = tvar \<^sort>\<open>numeral\<close>;
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    62
val numeral = cterm_of (Const (\<^const_name>\<open>numeral\<close>, \<^typ>\<open>num\<close> --> TVar numeral_tvar));
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    63
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    64
val uminus_tvar = tvar \<^sort>\<open>uminus\<close>;
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    65
val uminus = cterm_of (Const (\<^const_name>\<open>uminus\<close>, TVar uminus_tvar --> TVar uminus_tvar));
61150
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    66
74282
c2ee8d993d6a clarified signature: more scalable operations;
wenzelm
parents: 73020
diff changeset
    67
fun instT T v = Thm.instantiate_cterm (TVars.make [(v, T)], Vars.empty);
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    68
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    69
in
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    70
61150
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    71
fun mk_cnumber T 0 = instT T zero_tvar zero
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    72
  | mk_cnumber T 1 = instT T one_tvar one
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46497
diff changeset
    73
  | mk_cnumber T i =
61150
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    74
      if i > 0 then
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    75
        Thm.apply (instT T numeral_tvar numeral) (mk_cnumeral i)
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    76
      else
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    77
        Thm.apply (instT T uminus_tvar uminus)
d85d8f5e921b tuned -- avoid slightly odd @{cpat};
wenzelm
parents: 60642
diff changeset
    78
          (Thm.apply (instT T numeral_tvar numeral) (mk_cnumeral (~ i)));
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    79
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    80
end;
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
    81
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58399
diff changeset
    82
fun mk_number_syntax n =
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    83
  if n = 0 then Syntax.const \<^const_syntax>\<open>Groups.zero\<close>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    84
  else if n = 1 then Syntax.const \<^const_syntax>\<open>Groups.one\<close>
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    85
  else Syntax.const \<^const_syntax>\<open>numeral\<close> $ mk_num_syntax n;
58410
6d46ad54a2ab explicit separation of signed and unsigned numerals using existing lexical categories num and xnum
haftmann
parents: 58399
diff changeset
    86
25932
db0fd0ecdcd4 explicit auxiliary function for code setup
haftmann
parents: 25919
diff changeset
    87
db0fd0ecdcd4 explicit auxiliary function for code setup
haftmann
parents: 25919
diff changeset
    88
(* code generator *)
db0fd0ecdcd4 explicit auxiliary function for code setup
haftmann
parents: 25919
diff changeset
    89
28090
29af3c712d2b distributed literal code generation out of central infrastructure
haftmann
parents: 28064
diff changeset
    90
local open Basic_Code_Thingol in
29af3c712d2b distributed literal code generation out of central infrastructure
haftmann
parents: 28064
diff changeset
    91
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    92
fun dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\<open>Num.One\<close>, ... }) = SOME 1
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    93
  | dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\<open>Num.Bit0\<close>, ... } `$ t) =
68028
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 62597
diff changeset
    94
     (case dest_num_code t of
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    95
        SOME n => SOME (2 * n)
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    96
      | _ => NONE)
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 68028
diff changeset
    97
  | dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\<open>Num.Bit1\<close>, ... } `$ t) =
68028
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 62597
diff changeset
    98
     (case dest_num_code t of
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
    99
        SOME n => SOME (2 * n + 1)
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
   100
      | _ => NONE)
68028
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 62597
diff changeset
   101
  | dest_num_code _ = NONE;
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 61150
diff changeset
   102
58399
haftmann
parents: 55236
diff changeset
   103
fun add_code number_of preproc print target thy =
28090
29af3c712d2b distributed literal code generation out of central infrastructure
haftmann
parents: 28064
diff changeset
   104
  let
58399
haftmann
parents: 55236
diff changeset
   105
    fun pretty literals _ thm _ _ [(t, _)] =
68028
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 62597
diff changeset
   106
      case dest_num_code t of
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 62597
diff changeset
   107
        SOME n => (Code_Printer.str o print literals o preproc) n
1f9f973eed2a proper datatype for 8-bit characters
haftmann
parents: 62597
diff changeset
   108
      | NONE => Code_Printer.eqn_error thy thm "Illegal numeral expression: illegal term";
28090
29af3c712d2b distributed literal code generation out of central infrastructure
haftmann
parents: 28064
diff changeset
   109
  in
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51314
diff changeset
   110
    thy |> Code_Target.set_printings (Code_Symbol.Constant (number_of,
55148
7e1b7cb54114 avoid (now superfluous) indirect passing of constant names
haftmann
parents: 55147
diff changeset
   111
      [(target, SOME (Code_Printer.complex_const_syntax (1, pretty)))]))
28090
29af3c712d2b distributed literal code generation out of central infrastructure
haftmann
parents: 28064
diff changeset
   112
  end;
29af3c712d2b distributed literal code generation out of central infrastructure
haftmann
parents: 28064
diff changeset
   113
29af3c712d2b distributed literal code generation out of central infrastructure
haftmann
parents: 28064
diff changeset
   114
end; (*local*)
25932
db0fd0ecdcd4 explicit auxiliary function for code setup
haftmann
parents: 25919
diff changeset
   115
23575
543803006b3f Logical operations on numerals (see also HOL/hologic.ML).
wenzelm
parents:
diff changeset
   116
end;