author | nipkow |
Tue, 11 Oct 2022 10:45:42 +0200 | |
changeset 76259 | d1c26efb7a47 |
parent 75937 | 02b18f59f903 |
child 76387 | 8cb141384682 |
permissions | -rw-r--r-- |
23664 | 1 |
theory ComputeNumeral |
29804
e15b74577368
Added new Float theory and moved old Library/Float.thy to ComputeFloat
hoelzl
parents:
29668
diff
changeset
|
2 |
imports ComputeHOL ComputeFloat |
23664 | 3 |
begin |
4 |
||
5 |
(* equality for bit strings *) |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
6 |
lemmas biteq = eq_num_simps |
23664 | 7 |
|
8 |
(* x < y for bit strings *) |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
9 |
lemmas bitless = less_num_simps |
23664 | 10 |
|
11 |
(* x \<le> y for bit strings *) |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
12 |
lemmas bitle = le_num_simps |
23664 | 13 |
|
14 |
(* addition for bit strings *) |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
15 |
lemmas bitadd = add_num_simps |
23664 | 16 |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
60930
diff
changeset
|
17 |
(* multiplication for bit strings *) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
18 |
lemmas bitmul = mult_num_simps |
23664 | 19 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
20 |
lemmas bitarith = arith_simps |
23664 | 21 |
|
22 |
(* Normalization of nat literals *) |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
23 |
lemmas natnorm = one_eq_Numeral1_nat |
23664 | 24 |
|
25 |
fun natfac :: "nat \<Rightarrow> nat" |
|
38273 | 26 |
where "natfac n = (if n = 0 then 1 else n * (natfac (n - 1)))" |
23664 | 27 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
28 |
lemmas compute_natarith = |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
29 |
arith_simps rel_simps |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
30 |
diff_nat_numeral nat_numeral nat_0 nat_neg_numeral |
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
31 |
numeral_One [symmetric] |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
32 |
numeral_1_eq_Suc_0 [symmetric] |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
33 |
Suc_numeral natfac.simps |
23664 | 34 |
|
61694
6571c78c9667
Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents:
61609
diff
changeset
|
35 |
lemmas number_norm = numeral_One[symmetric] |
23664 | 36 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
37 |
lemmas compute_numberarith = |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
38 |
arith_simps rel_simps number_norm |
23664 | 39 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
40 |
lemmas compute_num_conversions = |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
60930
diff
changeset
|
41 |
of_nat_numeral of_nat_0 |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
42 |
nat_numeral nat_0 nat_neg_numeral |
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
60930
diff
changeset
|
43 |
of_int_numeral of_int_neg_numeral of_int_0 |
23664 | 44 |
|
62348 | 45 |
lemmas zpowerarith = power_numeral_even power_numeral_odd zpower_Pls int_pow_1 |
23664 | 46 |
|
47 |
(* div, mod *) |
|
48 |
||
60868
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
58953
diff
changeset
|
49 |
lemmas compute_div_mod = div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1 |
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
58953
diff
changeset
|
50 |
one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral |
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
58953
diff
changeset
|
51 |
one_div_minus_numeral one_mod_minus_numeral |
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
58953
diff
changeset
|
52 |
numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral |
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
58953
diff
changeset
|
53 |
numeral_div_minus_numeral numeral_mod_minus_numeral |
75937 | 54 |
div_minus_minus mod_minus_minus Euclidean_Division.adjust_div_eq of_bool_eq one_neq_zero |
60868
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
58953
diff
changeset
|
55 |
numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial |
75937 | 56 |
divmod_steps divmod_cancel divmod_step_def fst_conv snd_conv numeral_One |
57 |
case_prod_beta rel_simps Euclidean_Division.adjust_mod_def div_minus1_right mod_minus1_right |
|
60868
dd18c33c001e
direct bootstrap of integer division from natural division
haftmann
parents:
58953
diff
changeset
|
58 |
minus_minus numeral_times_numeral mult_zero_right mult_1_right |
23664 | 59 |
|
60 |
||
61 |
(* collecting all the theorems *) |
|
62 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
63 |
lemma even_0_int: "even (0::int) = True" |
23664 | 64 |
by simp |
65 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
66 |
lemma even_One_int: "even (numeral Num.One :: int) = False" |
23664 | 67 |
by simp |
68 |
||
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
69 |
lemma even_Bit0_int: "even (numeral (Num.Bit0 x) :: int) = True" |
58953 | 70 |
by (simp only: even_numeral) |
23664 | 71 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
72 |
lemma even_Bit1_int: "even (numeral (Num.Bit1 x) :: int) = False" |
58953 | 73 |
by (simp only: odd_numeral) |
23664 | 74 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46988
diff
changeset
|
75 |
lemmas compute_even = even_0_int even_One_int even_Bit0_int even_Bit1_int |
23664 | 76 |
|
61609
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents:
60930
diff
changeset
|
77 |
lemmas compute_numeral = compute_if compute_let compute_pair compute_bool |
23664 | 78 |
compute_natarith compute_numberarith max_def min_def compute_num_conversions zpowerarith compute_div_mod compute_even |
79 |
||
80 |
end |