| author | paulson <lp15@cam.ac.uk> | 
| Fri, 14 Jul 2023 14:21:25 +0100 | |
| changeset 78333 | c515bc9375b9 | 
| parent 77879 | dd222e2af01a | 
| child 82380 | ceb4f33d3073 | 
| permissions | -rw-r--r-- | 
| 28308 | 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 | 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: 
23575diff
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: 
58399diff
changeset | 10 | val mk_number_syntax: int -> term | 
| 62597 | 11 | val dest_num_syntax: term -> int | 
| 58399 | 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 | 20 | fun dest_num_syntax (Const (\<^const_syntax>\<open>Num.Bit0\<close>, _) $ t) = 2 * dest_num_syntax t | 
| 21 | | dest_num_syntax (Const (\<^const_syntax>\<open>Num.Bit1\<close>, _) $ t) = 2 * dest_num_syntax t + 1 | |
| 22 | | dest_num_syntax (Const (\<^const_syntax>\<open>Num.One\<close>, _)) = 1; | |
| 62597 | 23 | |
| 24 | fun mk_num_syntax n = | |
| 25 | if n > 0 then | |
| 73020 
b51515722274
tuned signature -- prefer Isabelle/ML structure Integer;
 wenzelm parents: 
69593diff
changeset | 26 | (case Integer.quot_rem n 2 of | 
| 69593 | 27 | (0, 1) => Syntax.const \<^const_syntax>\<open>One\<close> | 
| 28 | | (n, 0) => Syntax.const \<^const_syntax>\<open>Bit0\<close> $ mk_num_syntax n | |
| 29 | | (n, 1) => Syntax.const \<^const_syntax>\<open>Bit1\<close> $ mk_num_syntax n) | |
| 62597 | 30 | else raise Match | 
| 31 | ||
| 69593 | 32 | fun mk_cbit 0 = \<^cterm>\<open>Num.Bit0\<close> | 
| 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: 
46497diff
changeset | 36 | fun mk_cnumeral i = | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 37 | let | 
| 69593 | 38 | fun mk 1 = \<^cterm>\<open>Num.One\<close> | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 39 | | mk i = | 
| 24630 
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
 wenzelm parents: 
23575diff
changeset | 40 | let val (q, r) = Integer.div_mod i 2 in | 
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 41 | Thm.apply (mk_cbit r) (mk q) | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 42 | end | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 43 | in | 
| 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
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: 
46497diff
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 | 52 | val cterm_of = Thm.cterm_of \<^context>; | 
| 61150 | 53 | fun tvar S = (("'a", 0), S);
 | 
| 23575 
543803006b3f
Logical operations on numerals (see also HOL/hologic.ML).
 wenzelm parents: diff
changeset | 54 | |
| 69593 | 55 | val zero_tvar = tvar \<^sort>\<open>zero\<close>; | 
| 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 | 58 | val one_tvar = tvar \<^sort>\<open>one\<close>; | 
| 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: 
46497diff
changeset | 60 | |
| 69593 | 61 | val numeral_tvar = tvar \<^sort>\<open>numeral\<close>; | 
| 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 | 64 | val uminus_tvar = tvar \<^sort>\<open>uminus\<close>; | 
| 65 | val uminus = cterm_of (Const (\<^const_name>\<open>uminus\<close>, TVar uminus_tvar --> TVar uminus_tvar)); | |
| 61150 | 66 | |
| 77879 | 67 | fun instT T v = Thm.instantiate_cterm (TVars.make1 (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 | 71 | fun mk_cnumber T 0 = instT T zero_tvar zero | 
| 72 | | mk_cnumber T 1 = instT T one_tvar one | |
| 47108 
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
 huffman parents: 
46497diff
changeset | 73 | | mk_cnumber T i = | 
| 61150 | 74 | if i > 0 then | 
| 75 | Thm.apply (instT T numeral_tvar numeral) (mk_cnumeral i) | |
| 76 | else | |
| 77 | Thm.apply (instT T uminus_tvar uminus) | |
| 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: 
58399diff
changeset | 82 | fun mk_number_syntax n = | 
| 69593 | 83 | if n = 0 then Syntax.const \<^const_syntax>\<open>Groups.zero\<close> | 
| 84 | else if n = 1 then Syntax.const \<^const_syntax>\<open>Groups.one\<close> | |
| 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: 
58399diff
changeset | 86 | |
| 25932 | 87 | |
| 88 | (* code generator *) | |
| 89 | ||
| 28090 
29af3c712d2b
distributed literal code generation out of central infrastructure
 haftmann parents: 
28064diff
changeset | 90 | local open Basic_Code_Thingol in | 
| 
29af3c712d2b
distributed literal code generation out of central infrastructure
 haftmann parents: 
28064diff
changeset | 91 | |
| 69593 | 92 | fun dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\<open>Num.One\<close>, ... }) = SOME 1
 | 
| 93 |   | dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\<open>Num.Bit0\<close>, ... } `$ t) =
 | |
| 68028 | 94 | (case dest_num_code t of | 
| 62597 | 95 | SOME n => SOME (2 * n) | 
| 96 | | _ => NONE) | |
| 69593 | 97 |   | dest_num_code (IConst { sym = Code_Symbol.Constant \<^const_name>\<open>Num.Bit1\<close>, ... } `$ t) =
 | 
| 68028 | 98 | (case dest_num_code t of | 
| 62597 | 99 | SOME n => SOME (2 * n + 1) | 
| 100 | | _ => NONE) | |
| 68028 | 101 | | dest_num_code _ = NONE; | 
| 62597 | 102 | |
| 58399 | 103 | fun add_code number_of preproc print target thy = | 
| 28090 
29af3c712d2b
distributed literal code generation out of central infrastructure
 haftmann parents: 
28064diff
changeset | 104 | let | 
| 58399 | 105 | fun pretty literals _ thm _ _ [(t, _)] = | 
| 68028 | 106 | case dest_num_code t of | 
| 107 | SOME n => (Code_Printer.str o print literals o preproc) n | |
| 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: 
28064diff
changeset | 109 | in | 
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51314diff
changeset | 110 | thy |> Code_Target.set_printings (Code_Symbol.Constant (number_of, | 
| 55148 
7e1b7cb54114
avoid (now superfluous) indirect passing of constant names
 haftmann parents: 
55147diff
changeset | 111 | [(target, SOME (Code_Printer.complex_const_syntax (1, pretty)))])) | 
| 28090 
29af3c712d2b
distributed literal code generation out of central infrastructure
 haftmann parents: 
28064diff
changeset | 112 | end; | 
| 
29af3c712d2b
distributed literal code generation out of central infrastructure
 haftmann parents: 
28064diff
changeset | 113 | |
| 
29af3c712d2b
distributed literal code generation out of central infrastructure
 haftmann parents: 
28064diff
changeset | 114 | end; (*local*) | 
| 25932 | 115 | |
| 23575 
543803006b3f
Logical operations on numerals (see also HOL/hologic.ML).
 wenzelm parents: diff
changeset | 116 | end; |