src/HOL/Parity.thy
author paulson <lp15@cam.ac.uk>
Tue, 16 Jan 2024 13:40:09 +0000
changeset 79492 c1b0f64eb865
parent 79118 486a32079c60
child 79531 22a137a6de44
permissions -rw-r--r--
A few new results (mostly brought in from other developments)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 36840
diff changeset
     1
(*  Title:      HOL/Parity.thy
b460124855b8 tuned headers;
wenzelm
parents: 36840
diff changeset
     2
    Author:     Jeremy Avigad
b460124855b8 tuned headers;
wenzelm
parents: 36840
diff changeset
     3
    Author:     Jacques D. Fleuriot
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     4
*)
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     5
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60562
diff changeset
     6
section \<open>Parity in rings and semirings\<close>
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     7
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
     8
theory Parity
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 76387
diff changeset
     9
  imports Euclidean_Rings
21256
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    10
begin
47195501ecf7 moved theories Parity, GCD, Binomial to Library;
wenzelm
parents:
diff changeset
    11
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61531
diff changeset
    12
subsection \<open>Ring structures with parity and \<open>even\<close>/\<open>odd\<close> predicates\<close>
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    13
70341
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
    14
class semiring_parity = comm_semiring_1 + semiring_modulo +
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    15
  assumes mod_2_eq_odd: \<open>a mod 2 = of_bool (\<not> 2 dvd a)\<close>
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    16
    and odd_one [simp]: \<open>\<not> 2 dvd 1\<close>
66839
909ba5ed93dd clarified parity
haftmann
parents: 66816
diff changeset
    17
begin
909ba5ed93dd clarified parity
haftmann
parents: 66816
diff changeset
    18
58740
cb9d84d3e7f2 turn even into an abbreviation
haftmann
parents: 58718
diff changeset
    19
abbreviation even :: "'a \<Rightarrow> bool"
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    20
  where \<open>even a \<equiv> 2 dvd a\<close>
54228
229282d53781 purely algebraic foundation for even/odd
haftmann
parents: 54227
diff changeset
    21
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    22
abbreviation odd :: "'a \<Rightarrow> bool"
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    23
  where \<open>odd a \<equiv> \<not> 2 dvd a\<close>
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
    24
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    25
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    26
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    27
class ring_parity = ring + semiring_parity
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    28
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    29
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    30
subclass comm_ring_1 ..
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    31
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    32
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    33
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    34
instance nat :: semiring_parity
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    35
  by standard (simp_all add: dvd_eq_mod_eq_0)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    36
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    37
instance int :: ring_parity
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    38
  by standard (auto simp add: dvd_eq_mod_eq_0)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    39
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    40
context semiring_parity
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    41
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
    42
58690
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
    43
lemma evenE [elim?]:
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    44
  assumes \<open>even a\<close>
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    45
  obtains b where \<open>a = 2 * b\<close>
58740
cb9d84d3e7f2 turn even into an abbreviation
haftmann
parents: 58718
diff changeset
    46
  using assms by (rule dvdE)
58690
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
    47
58681
a478a0742a8e legacy cleanup
haftmann
parents: 58680
diff changeset
    48
lemma oddE [elim?]:
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    49
  assumes \<open>odd a\<close>
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    50
  obtains b where \<open>a = 2 * b + 1\<close>
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
    51
proof -
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    52
  have \<open>a = 2 * (a div 2) + a mod 2\<close>
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
    53
    by (simp add: mult_div_mod_eq)
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    54
  with assms have \<open>a = 2 * (a div 2) + 1\<close>
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    55
    by (simp add: mod_2_eq_odd)
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    56
  then show thesis ..
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
    57
qed
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
    58
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
    59
lemma of_bool_odd_eq_mod_2:
79117
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    60
  \<open>of_bool (odd a) = a mod 2\<close>
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    61
  by (simp add: mod_2_eq_odd)
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    62
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    63
lemma odd_of_bool_self [simp]:
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    64
  \<open>odd (of_bool p) \<longleftrightarrow> p\<close>
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    65
  by (cases p) simp_all
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    66
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    67
lemma not_mod_2_eq_0_eq_1 [simp]:
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    68
  \<open>a mod 2 \<noteq> 0 \<longleftrightarrow> a mod 2 = 1\<close>
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    69
  by (simp add: mod_2_eq_odd)
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    70
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    71
lemma not_mod_2_eq_1_eq_0 [simp]:
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    72
  \<open>a mod 2 \<noteq> 1 \<longleftrightarrow> a mod 2 = 0\<close>
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
    73
  by (simp add: mod_2_eq_odd)
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
    74
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    75
lemma even_iff_mod_2_eq_zero:
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    76
  \<open>2 dvd a \<longleftrightarrow> a mod 2 = 0\<close>
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    77
  by (simp add: mod_2_eq_odd)
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    78
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    79
lemma odd_iff_mod_2_eq_one:
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    80
  \<open>\<not> 2 dvd a \<longleftrightarrow> a mod 2 = 1\<close>
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    81
  by (simp add: mod_2_eq_odd)
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
    82
71426
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    83
lemma even_mod_2_iff [simp]:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    84
  \<open>even (a mod 2) \<longleftrightarrow> even a\<close>
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    85
  by (simp add: mod_2_eq_odd)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    86
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    87
lemma mod2_eq_if:
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    88
  "a mod 2 = (if even a then 0 else 1)"
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    89
  by (simp add: mod_2_eq_odd)
745e518d3d0b easy abstraction over pointwise bit operations
haftmann
parents: 71424
diff changeset
    90
79117
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    91
lemma zero_mod_two_eq_zero [simp]:
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    92
  \<open>0 mod 2 = 0\<close>
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    93
  by (simp add: mod_2_eq_odd)
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    94
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    95
lemma one_mod_two_eq_one [simp]:
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    96
  \<open>1 mod 2 = 1\<close>
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    97
  by (simp add: mod_2_eq_odd)
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    98
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
    99
lemma parity_cases [case_names even odd]:
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
   100
  assumes \<open>even a \<Longrightarrow> a mod 2 = 0 \<Longrightarrow> P\<close>
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
   101
  assumes \<open>odd a \<Longrightarrow> a mod 2 = 1 \<Longrightarrow> P\<close>
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
   102
  shows P
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
   103
  using assms by (auto simp add: mod_2_eq_odd)
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
   104
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   105
lemma even_zero [simp]:
79117
7476818dfd5d generalized
haftmann
parents: 78937
diff changeset
   106
  \<open>even 0\<close>
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   107
  by (fact dvd_0_right)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   108
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   109
lemma odd_even_add:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   110
  "even (a + b)" if "odd a" and "odd b"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   111
proof -
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   112
  from that obtain c d where "a = 2 * c + 1" and "b = 2 * d + 1"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   113
    by (blast elim: oddE)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   114
  then have "a + b = 2 * c + 2 * d + (1 + 1)"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   115
    by (simp only: ac_simps)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   116
  also have "\<dots> = 2 * (c + d + 1)"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   117
    by (simp add: algebra_simps)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   118
  finally show ?thesis ..
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   119
qed
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   120
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   121
lemma even_add [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   122
  "even (a + b) \<longleftrightarrow> (even a \<longleftrightarrow> even b)"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   123
  by (auto simp add: dvd_add_right_iff dvd_add_left_iff odd_even_add)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   124
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   125
lemma odd_add [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   126
  "odd (a + b) \<longleftrightarrow> \<not> (odd a \<longleftrightarrow> odd b)"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   127
  by simp
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   128
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   129
lemma even_plus_one_iff [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   130
  "even (a + 1) \<longleftrightarrow> odd a"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   131
  by (auto simp add: dvd_add_right_iff intro: odd_even_add)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   132
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   133
lemma even_mult_iff [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   134
  "even (a * b) \<longleftrightarrow> even a \<or> even b" (is "?P \<longleftrightarrow> ?Q")
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   135
proof
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   136
  assume ?Q
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   137
  then show ?P
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   138
    by auto
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   139
next
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   140
  assume ?P
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   141
  show ?Q
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   142
  proof (rule ccontr)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   143
    assume "\<not> (even a \<or> even b)"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   144
    then have "odd a" and "odd b"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   145
      by auto
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   146
    then obtain r s where "a = 2 * r + 1" and "b = 2 * s + 1"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   147
      by (blast elim: oddE)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   148
    then have "a * b = (2 * r + 1) * (2 * s + 1)"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   149
      by simp
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   150
    also have "\<dots> = 2 * (2 * r * s + r + s) + 1"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   151
      by (simp add: algebra_simps)
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   152
    finally have "odd (a * b)"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   153
      by simp
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   154
    with \<open>?P\<close> show False
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   155
      by auto
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   156
  qed
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   157
qed
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   158
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   159
lemma even_numeral [simp]: "even (numeral (Num.Bit0 n))"
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   160
proof -
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   161
  have "even (2 * numeral n)"
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   162
    unfolding even_mult_iff by simp
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   163
  then have "even (numeral n + numeral n)"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   164
    unfolding mult_2 .
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   165
  then show ?thesis
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   166
    unfolding numeral.simps .
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   167
qed
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   168
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   169
lemma odd_numeral [simp]: "odd (numeral (Num.Bit1 n))"
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   170
proof
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   171
  assume "even (numeral (num.Bit1 n))"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   172
  then have "even (numeral n + numeral n + 1)"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   173
    unfolding numeral.simps .
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   174
  then have "even (2 * numeral n + 1)"
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   175
    unfolding mult_2 .
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   176
  then have "2 dvd numeral n * 2 + 1"
58740
cb9d84d3e7f2 turn even into an abbreviation
haftmann
parents: 58718
diff changeset
   177
    by (simp add: ac_simps)
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   178
  then have "2 dvd 1"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   179
    using dvd_add_times_triv_left_iff [of 2 "numeral n" 1] by simp
58678
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   180
  then show False by simp
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   181
qed
398e05aa84d4 purely algebraic characterization of even and odd
haftmann
parents: 58645
diff changeset
   182
71755
318695613bb7 more complete rules on numerals
haftmann
parents: 71535
diff changeset
   183
lemma odd_numeral_BitM [simp]:
318695613bb7 more complete rules on numerals
haftmann
parents: 71535
diff changeset
   184
  \<open>odd (numeral (Num.BitM w))\<close>
318695613bb7 more complete rules on numerals
haftmann
parents: 71535
diff changeset
   185
  by (cases w) simp_all
318695613bb7 more complete rules on numerals
haftmann
parents: 71535
diff changeset
   186
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   187
lemma even_power [simp]: "even (a ^ n) \<longleftrightarrow> even a \<and> n > 0"
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   188
  by (induct n) auto
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   189
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   190
lemma even_prod_iff:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   191
  \<open>even (prod f A) \<longleftrightarrow> (\<exists>a\<in>A. even (f a))\<close> if \<open>finite A\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   192
  using that by (induction A) simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   193
71412
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   194
lemma mask_eq_sum_exp:
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   195
  \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   196
proof -
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   197
  have *: \<open>{q. q < Suc m} = insert m {q. q < m}\<close> for m
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   198
    by auto
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   199
  have \<open>2 ^ n = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1\<close>
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   200
    by (induction n) (simp_all add: ac_simps mult_2 *)
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   201
  then have \<open>2 ^ n - 1 = (\<Sum>m\<in>{q. q < n}. 2 ^ m) + 1 - 1\<close>
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   202
    by simp
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   203
  then show ?thesis
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   204
    by simp
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   205
qed
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   206
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   207
lemma (in -) mask_eq_sum_exp_nat:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   208
  \<open>2 ^ n - Suc 0 = (\<Sum>m\<in>{q. q < n}. 2 ^ m)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   209
  using mask_eq_sum_exp [where ?'a = nat] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   210
70341
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   211
end
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   212
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   213
context ring_parity
70341
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   214
begin
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   215
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   216
lemma even_minus:
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   217
  "even (- a) \<longleftrightarrow> even a"
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   218
  by (fact dvd_minus_iff)
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   219
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   220
lemma even_diff [simp]:
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   221
  "even (a - b) \<longleftrightarrow> even (a + b)"
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   222
  using even_add [of a "- b"] by simp
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   223
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   224
end
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   225
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   226
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69502
diff changeset
   227
subsection \<open>Instance for \<^typ>\<open>nat\<close>\<close>
66808
1907167b6038 elementary definition of division on natural numbers
haftmann
parents: 66582
diff changeset
   228
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   229
lemma even_Suc_Suc_iff [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   230
  "even (Suc (Suc n)) \<longleftrightarrow> even n"
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   231
  using dvd_add_triv_right_iff [of 2 n] by simp
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   232
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   233
lemma even_Suc [simp]: "even (Suc n) \<longleftrightarrow> odd n"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   234
  using even_plus_one_iff [of n] by simp
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   235
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   236
lemma even_diff_nat [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   237
  "even (m - n) \<longleftrightarrow> m < n \<or> even (m + n)" for m n :: nat
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   238
proof (cases "n \<le> m")
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   239
  case True
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   240
  then have "m - n + n * 2 = m + n" by (simp add: mult_2_right)
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   241
  moreover have "even (m - n) \<longleftrightarrow> even (m - n + n * 2)" by simp
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   242
  ultimately have "even (m - n) \<longleftrightarrow> even (m + n)" by (simp only:)
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   243
  then show ?thesis by auto
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   244
next
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   245
  case False
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   246
  then show ?thesis by simp
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   247
qed
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   248
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   249
lemma odd_pos:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   250
  "odd n \<Longrightarrow> 0 < n" for n :: nat
58690
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
   251
  by (auto elim: oddE)
60343
063698416239 correct sort constraints for abbreviations in type classes
haftmann
parents: 59816
diff changeset
   252
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   253
lemma Suc_double_not_eq_double:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   254
  "Suc (2 * m) \<noteq> 2 * n"
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   255
proof
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   256
  assume "Suc (2 * m) = 2 * n"
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   257
  moreover have "odd (Suc (2 * m))" and "even (2 * n)"
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   258
    by simp_all
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   259
  ultimately show False by simp
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   260
qed
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   261
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   262
lemma double_not_eq_Suc_double:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   263
  "2 * m \<noteq> Suc (2 * n)"
62597
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   264
  using Suc_double_not_eq_double [of n m] by simp
b3f2b8c906a6 model characters directly as range 0..255
haftmann
parents: 62083
diff changeset
   265
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   266
lemma odd_Suc_minus_one [simp]: "odd n \<Longrightarrow> Suc (n - Suc 0) = n"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   267
  by (auto elim: oddE)
60343
063698416239 correct sort constraints for abbreviations in type classes
haftmann
parents: 59816
diff changeset
   268
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   269
lemma even_Suc_div_two [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   270
  "even n \<Longrightarrow> Suc n div 2 = n div 2"
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   271
  by auto
60343
063698416239 correct sort constraints for abbreviations in type classes
haftmann
parents: 59816
diff changeset
   272
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   273
lemma odd_Suc_div_two [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   274
  "odd n \<Longrightarrow> Suc n div 2 = Suc (n div 2)"
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   275
  by (auto elim: oddE)
60343
063698416239 correct sort constraints for abbreviations in type classes
haftmann
parents: 59816
diff changeset
   276
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   277
lemma odd_two_times_div_two_nat [simp]:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   278
  assumes "odd n"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   279
  shows "2 * (n div 2) = n - (1 :: nat)"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   280
proof -
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   281
  from assms have "2 * (n div 2) + 1 = n"
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   282
    by (auto elim: oddE)
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   283
  then have "Suc (2 * (n div 2)) - 1 = n - 1"
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   284
    by simp
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   285
  then show ?thesis
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   286
    by simp
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   287
qed
58680
6b2fa479945f more algebraic deductions for facts on even/odd
haftmann
parents: 58679
diff changeset
   288
70341
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   289
lemma not_mod2_eq_Suc_0_eq_0 [simp]:
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   290
  "n mod 2 \<noteq> Suc 0 \<longleftrightarrow> n mod 2 = 0"
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   291
  using not_mod_2_eq_1_eq_0 [of n] by simp
972c0c744e7c generalized type classes for parity to cover word types also, which contain zero divisors
haftmann
parents: 70340
diff changeset
   292
69502
0cf906072e20 more rules
haftmann
parents: 69198
diff changeset
   293
lemma odd_card_imp_not_empty:
0cf906072e20 more rules
haftmann
parents: 69198
diff changeset
   294
  \<open>A \<noteq> {}\<close> if \<open>odd (card A)\<close>
0cf906072e20 more rules
haftmann
parents: 69198
diff changeset
   295
  using that by auto
0cf906072e20 more rules
haftmann
parents: 69198
diff changeset
   296
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   297
lemma nat_induct2 [case_names 0 1 step]:
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   298
  assumes "P 0" "P 1" and step: "\<And>n::nat. P n \<Longrightarrow> P (n + 2)"
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   299
  shows "P n"
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   300
proof (induct n rule: less_induct)
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   301
  case (less n)
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   302
  show ?case
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   303
  proof (cases "n < Suc (Suc 0)")
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   304
    case True
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   305
    then show ?thesis
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   306
      using assms by (auto simp: less_Suc_eq)
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   307
  next
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   308
    case False
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   309
    then obtain k where k: "n = Suc (Suc k)"
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   310
      by (force simp: not_less nat_le_iff_add)
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   311
    then have "k<n"
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   312
      by simp
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   313
    with less assms have "P (k+2)"
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   314
      by blast
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   315
    then show ?thesis
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   316
      by (simp add: k)
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   317
  qed
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70353
diff changeset
   318
qed
58687
5469874b0228 even more cleanup
haftmann
parents: 58681
diff changeset
   319
78668
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   320
lemma mod_double_nat:
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   321
  \<open>n mod (2 * m) = n mod m \<or> n mod (2 * m) = n mod m + m\<close>
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   322
  for m n :: nat
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   323
  by (cases \<open>even (n div m)\<close>)
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   324
    (simp_all add: mod_mult2_eq ac_simps even_iff_mod_2_eq_zero)
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   325
71412
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   326
context semiring_parity
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   327
begin
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   328
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   329
lemma even_sum_iff:
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   330
  \<open>even (sum f A) \<longleftrightarrow> even (card {a\<in>A. odd (f a)})\<close> if \<open>finite A\<close>
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   331
using that proof (induction A)
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   332
  case empty
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   333
  then show ?case
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   334
    by simp
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   335
next
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   336
  case (insert a A)
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   337
  moreover have \<open>{b \<in> insert a A. odd (f b)} = (if odd (f a) then {a} else {}) \<union> {b \<in> A. odd (f b)}\<close>
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   338
    by auto
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   339
  ultimately show ?case
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   340
    by simp
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   341
qed
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   342
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   343
lemma even_mask_iff [simp]:
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   344
  \<open>even (2 ^ n - 1) \<longleftrightarrow> n = 0\<close>
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   345
proof (cases \<open>n = 0\<close>)
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   346
  case True
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   347
  then show ?thesis
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   348
    by simp
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   349
next
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   350
  case False
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   351
  then have \<open>{a. a = 0 \<and> a < n} = {0}\<close>
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   352
    by auto
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   353
  then show ?thesis
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   354
    by (auto simp add: mask_eq_sum_exp even_sum_iff)
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   355
qed
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   356
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   357
lemma even_of_nat_iff [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   358
  "even (of_nat n) \<longleftrightarrow> even n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   359
  by (induction n) simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   360
71412
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   361
end
96d126844adc more theorems
haftmann
parents: 71408
diff changeset
   362
71138
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71094
diff changeset
   363
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60562
diff changeset
   364
subsection \<open>Parity and powers\<close>
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   365
61531
ab2e862263e7 Rounding function, uniform limits, cotangent, binomial identities
eberlm
parents: 60867
diff changeset
   366
context ring_1
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   367
begin
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   368
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   369
lemma power_minus_even [simp]: "even n \<Longrightarrow> (- a) ^ n = a ^ n"
58690
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
   370
  by (auto elim: evenE)
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   371
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   372
lemma power_minus_odd [simp]: "odd n \<Longrightarrow> (- a) ^ n = - (a ^ n)"
58690
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
   373
  by (auto elim: oddE)
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
   374
66815
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   375
lemma uminus_power_if:
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   376
  "(- a) ^ n = (if even n then a ^ n else - (a ^ n))"
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   377
  by auto
93c6632ddf44 one uniform type class for parity structures
haftmann
parents: 66808
diff changeset
   378
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   379
lemma neg_one_even_power [simp]: "even n \<Longrightarrow> (- 1) ^ n = 1"
58690
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
   380
  by simp
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   381
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   382
lemma neg_one_odd_power [simp]: "odd n \<Longrightarrow> (- 1) ^ n = - 1"
58690
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
   383
  by simp
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   384
66582
2b49d4888cb8 another fact on (- 1) ^ _
bulwahn
parents: 64785
diff changeset
   385
lemma neg_one_power_add_eq_neg_one_power_diff: "k \<le> n \<Longrightarrow> (- 1) ^ (n + k) = (- 1) ^ (n - k)"
2b49d4888cb8 another fact on (- 1) ^ _
bulwahn
parents: 64785
diff changeset
   386
  by (cases "even (n + k)") auto
2b49d4888cb8 another fact on (- 1) ^ _
bulwahn
parents: 64785
diff changeset
   387
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67083
diff changeset
   388
lemma minus_one_power_iff: "(- 1) ^ n = (if even n then 1 else - 1)"
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67083
diff changeset
   389
  by (induct n) auto
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67083
diff changeset
   390
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   391
end
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   392
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   393
context linordered_idom
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   394
begin
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   395
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   396
lemma zero_le_even_power: "even n \<Longrightarrow> 0 \<le> a ^ n"
58690
5c5c14844738 standard elimination rule for even
haftmann
parents: 58689
diff changeset
   397
  by (auto elim: evenE)
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   398
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   399
lemma zero_le_odd_power: "odd n \<Longrightarrow> 0 \<le> a ^ n \<longleftrightarrow> 0 \<le> a"
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   400
  by (auto simp add: power_even_eq zero_le_mult_iff elim: oddE)
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   401
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   402
lemma zero_le_power_eq: "0 \<le> a ^ n \<longleftrightarrow> even n \<or> odd n \<and> 0 \<le> a"
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   403
  by (auto simp add: zero_le_even_power zero_le_odd_power)
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   404
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   405
lemma zero_less_power_eq: "0 < a ^ n \<longleftrightarrow> n = 0 \<or> even n \<and> a \<noteq> 0 \<or> odd n \<and> 0 < a"
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   406
proof -
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   407
  have [simp]: "0 = a ^ n \<longleftrightarrow> a = 0 \<and> n > 0"
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 58778
diff changeset
   408
    unfolding power_eq_0_iff [of a n, symmetric] by blast
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   409
  show ?thesis
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   410
    unfolding less_le zero_le_power_eq by auto
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   411
qed
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   412
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   413
lemma power_less_zero_eq [simp]: "a ^ n < 0 \<longleftrightarrow> odd n \<and> a < 0"
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   414
  unfolding not_le [symmetric] zero_le_power_eq by auto
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   415
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   416
lemma power_le_zero_eq: "a ^ n \<le> 0 \<longleftrightarrow> n > 0 \<and> (odd n \<and> a \<le> 0 \<or> even n \<and> a = 0)"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   417
  unfolding not_less [symmetric] zero_less_power_eq by auto
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   418
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   419
lemma power_even_abs: "even n \<Longrightarrow> \<bar>a\<bar> ^ n = a ^ n"
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   420
  using power_abs [of a n] by (simp add: zero_le_even_power)
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   421
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   422
lemma power_mono_even:
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   423
  assumes "even n" and "\<bar>a\<bar> \<le> \<bar>b\<bar>"
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   424
  shows "a ^ n \<le> b ^ n"
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   425
proof -
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   426
  have "0 \<le> \<bar>a\<bar>" by auto
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   427
  with \<open>\<bar>a\<bar> \<le> \<bar>b\<bar>\<close> have "\<bar>a\<bar> ^ n \<le> \<bar>b\<bar> ^ n"
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   428
    by (rule power_mono)
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   429
  with \<open>even n\<close> show ?thesis
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   430
    by (simp add: power_even_abs)
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   431
qed
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   432
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   433
lemma power_mono_odd:
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   434
  assumes "odd n" and "a \<le> b"
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   435
  shows "a ^ n \<le> b ^ n"
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   436
proof (cases "b < 0")
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   437
  case True
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   438
  with \<open>a \<le> b\<close> have "- b \<le> - a" and "0 \<le> - b" by auto
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   439
  then have "(- b) ^ n \<le> (- a) ^ n" by (rule power_mono)
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60562
diff changeset
   440
  with \<open>odd n\<close> show ?thesis by simp
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   441
next
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   442
  case False
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   443
  then have "0 \<le> b" by auto
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   444
  show ?thesis
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   445
  proof (cases "a < 0")
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   446
    case True
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   447
    then have "n \<noteq> 0" and "a \<le> 0" using \<open>odd n\<close> [THEN odd_pos] by auto
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60562
diff changeset
   448
    then have "a ^ n \<le> 0" unfolding power_le_zero_eq using \<open>odd n\<close> by auto
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   449
    moreover from \<open>0 \<le> b\<close> have "0 \<le> b ^ n" by auto
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   450
    ultimately show ?thesis by auto
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   451
  next
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   452
    case False
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   453
    then have "0 \<le> a" by auto
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   454
    with \<open>a \<le> b\<close> show ?thesis
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   455
      using power_mono by auto
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   456
  qed
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   457
qed
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61799
diff changeset
   458
60758
d8d85a8172b5 isabelle update_cartouches;
wenzelm
parents: 60562
diff changeset
   459
text \<open>Simplify, when the exponent is a numeral\<close>
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   460
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   461
lemma zero_le_power_eq_numeral [simp]:
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   462
  "0 \<le> a ^ numeral w \<longleftrightarrow> even (numeral w :: nat) \<or> odd (numeral w :: nat) \<and> 0 \<le> a"
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   463
  by (fact zero_le_power_eq)
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   464
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   465
lemma zero_less_power_eq_numeral [simp]:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   466
  "0 < a ^ numeral w \<longleftrightarrow>
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   467
    numeral w = (0 :: nat) \<or>
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   468
    even (numeral w :: nat) \<and> a \<noteq> 0 \<or>
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   469
    odd (numeral w :: nat) \<and> 0 < a"
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   470
  by (fact zero_less_power_eq)
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   471
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   472
lemma power_le_zero_eq_numeral [simp]:
63654
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   473
  "a ^ numeral w \<le> 0 \<longleftrightarrow>
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   474
    (0 :: nat) < numeral w \<and>
f90e3926e627 misc tuning and modernization;
wenzelm
parents: 62597
diff changeset
   475
    (odd (numeral w :: nat) \<and> a \<le> 0 \<or> even (numeral w :: nat) \<and> a = 0)"
58689
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   476
  by (fact power_le_zero_eq)
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   477
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   478
lemma power_less_zero_eq_numeral [simp]:
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   479
  "a ^ numeral w < 0 \<longleftrightarrow> odd (numeral w :: nat) \<and> a < 0"
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   480
  by (fact power_less_zero_eq)
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   481
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   482
lemma power_even_abs_numeral [simp]:
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   483
  "even (numeral w :: nat) \<Longrightarrow> \<bar>a\<bar> ^ numeral w = a ^ numeral w"
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   484
  by (fact power_even_abs)
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   485
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   486
end
ee5bf401cfa7 tuned facts on even and power
haftmann
parents: 58688
diff changeset
   487
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66815
diff changeset
   488
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 69502
diff changeset
   489
subsection \<open>Instance for \<^typ>\<open>int\<close>\<close>
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   490
  
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
   491
lemma even_diff_iff:
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66815
diff changeset
   492
  "even (k - l) \<longleftrightarrow> even (k + l)" for k l :: int
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
   493
  by (fact even_diff)
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66815
diff changeset
   494
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
   495
lemma even_abs_add_iff:
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66815
diff changeset
   496
  "even (\<bar>k\<bar> + l) \<longleftrightarrow> even (k + l)" for k l :: int
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
   497
  by simp
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66815
diff changeset
   498
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
   499
lemma even_add_abs_iff:
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66815
diff changeset
   500
  "even (k + \<bar>l\<bar>) \<longleftrightarrow> even (k + l)" for k l :: int
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
   501
  by simp
66816
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66815
diff changeset
   502
212a3334e7da more fundamental definition of div and mod on int
haftmann
parents: 66815
diff changeset
   503
lemma even_nat_iff: "0 \<le> k \<Longrightarrow> even (nat k) \<longleftrightarrow> even k"
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74101
diff changeset
   504
  by (simp add: even_of_nat_iff [of "nat k", where ?'a = int, symmetric])
71138
9de7f1067520 strengthened type class for bit operations
haftmann
parents: 71094
diff changeset
   505
71837
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   506
context
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   507
  assumes "SORT_CONSTRAINT('a::division_ring)"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   508
begin
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   509
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   510
lemma power_int_minus_left:
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   511
  "power_int (-a :: 'a) n = (if even n then power_int a n else -power_int a n)"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   512
  by (auto simp: power_int_def minus_one_power_iff even_nat_iff)
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   513
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   514
lemma power_int_minus_left_even [simp]: "even n \<Longrightarrow> power_int (-a :: 'a) n = power_int a n"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   515
  by (simp add: power_int_minus_left)
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   516
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   517
lemma power_int_minus_left_odd [simp]: "odd n \<Longrightarrow> power_int (-a :: 'a) n = -power_int a n"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   518
  by (simp add: power_int_minus_left)
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   519
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   520
lemma power_int_minus_left_distrib:
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   521
  "NO_MATCH (-1) x \<Longrightarrow> power_int (-a :: 'a) n = power_int (-1) n * power_int a n"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   522
  by (simp add: power_int_minus_left)
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   523
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   524
lemma power_int_minus_one_minus: "power_int (-1 :: 'a) (-n) = power_int (-1) n"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   525
  by (simp add: power_int_minus_left)
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   526
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   527
lemma power_int_minus_one_diff_commute: "power_int (-1 :: 'a) (a - b) = power_int (-1) (b - a)"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   528
  by (subst power_int_minus_one_minus [symmetric]) auto
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   529
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   530
lemma power_int_minus_one_mult_self [simp]:
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   531
  "power_int (-1 :: 'a) m * power_int (-1) m = 1"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   532
  by (simp add: power_int_minus_left)
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   533
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   534
lemma power_int_minus_one_mult_self' [simp]:
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   535
  "power_int (-1 :: 'a) m * (power_int (-1) m * b) = b"
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   536
  by (simp add: power_int_minus_left)
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   537
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   538
end
dca11678c495 new constant power_int in HOL
Manuel Eberl <eberlm@in.tum.de>
parents: 71822
diff changeset
   539
75937
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
   540
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   541
subsection \<open>Special case: euclidean rings structurally containing the natural numbers\<close>
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   542
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   543
class linordered_euclidean_semiring = discrete_linordered_semidom + unique_euclidean_semiring +
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   544
  assumes of_nat_div: "of_nat (m div n) = of_nat m div of_nat n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   545
    and division_segment_of_nat [simp]: "division_segment (of_nat n) = 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   546
    and division_segment_euclidean_size [simp]: "division_segment a * of_nat (euclidean_size a) = a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   547
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   548
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   549
lemma division_segment_eq_iff:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   550
  "a = b" if "division_segment a = division_segment b"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   551
    and "euclidean_size a = euclidean_size b"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   552
  using that division_segment_euclidean_size [of a] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   553
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   554
lemma euclidean_size_of_nat [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   555
  "euclidean_size (of_nat n) = n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   556
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   557
  have "division_segment (of_nat n) * of_nat (euclidean_size (of_nat n)) = of_nat n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   558
    by (fact division_segment_euclidean_size)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   559
  then show ?thesis by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   560
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   561
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   562
lemma of_nat_euclidean_size:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   563
  "of_nat (euclidean_size a) = a div division_segment a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   564
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   565
  have "of_nat (euclidean_size a) = division_segment a * of_nat (euclidean_size a) div division_segment a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   566
    by (subst nonzero_mult_div_cancel_left) simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   567
  also have "\<dots> = a div division_segment a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   568
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   569
  finally show ?thesis .
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   570
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   571
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   572
lemma division_segment_1 [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   573
  "division_segment 1 = 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   574
  using division_segment_of_nat [of 1] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   575
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   576
lemma division_segment_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   577
  "division_segment (numeral k) = 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   578
  using division_segment_of_nat [of "numeral k"] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   579
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   580
lemma euclidean_size_1 [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   581
  "euclidean_size 1 = 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   582
  using euclidean_size_of_nat [of 1] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   583
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   584
lemma euclidean_size_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   585
  "euclidean_size (numeral k) = numeral k"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   586
  using euclidean_size_of_nat [of "numeral k"] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   587
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   588
lemma of_nat_dvd_iff:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   589
  "of_nat m dvd of_nat n \<longleftrightarrow> m dvd n" (is "?P \<longleftrightarrow> ?Q")
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   590
proof (cases "m = 0")
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   591
  case True
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   592
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   593
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   594
next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   595
  case False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   596
  show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   597
  proof
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   598
    assume ?Q
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   599
    then show ?P
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   600
      by auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   601
  next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   602
    assume ?P
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   603
    with False have "of_nat n = of_nat n div of_nat m * of_nat m"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   604
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   605
    then have "of_nat n = of_nat (n div m * m)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   606
      by (simp add: of_nat_div)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   607
    then have "n = n div m * m"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   608
      by (simp only: of_nat_eq_iff)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   609
    then have "n = m * (n div m)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   610
      by (simp add: ac_simps)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   611
    then show ?Q ..
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   612
  qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   613
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   614
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   615
lemma of_nat_mod:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   616
  "of_nat (m mod n) = of_nat m mod of_nat n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   617
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   618
  have "of_nat m div of_nat n * of_nat n + of_nat m mod of_nat n = of_nat m"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   619
    by (simp add: div_mult_mod_eq)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   620
  also have "of_nat m = of_nat (m div n * n + m mod n)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   621
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   622
  finally show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   623
    by (simp only: of_nat_div of_nat_mult of_nat_add) simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   624
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   625
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   626
lemma one_div_two_eq_zero [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   627
  "1 div 2 = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   628
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   629
  from of_nat_div [symmetric] have "of_nat 1 div of_nat 2 = of_nat 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   630
    by (simp only:) simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   631
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   632
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   633
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   634
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   635
lemma one_mod_2_pow_eq [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   636
  "1 mod (2 ^ n) = of_bool (n > 0)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   637
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   638
  have "1 mod (2 ^ n) = of_nat (1 mod (2 ^ n))"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   639
    using of_nat_mod [of 1 "2 ^ n"] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   640
  also have "\<dots> = of_bool (n > 0)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   641
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   642
  finally show ?thesis .
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   643
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   644
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   645
lemma one_div_2_pow_eq [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   646
  "1 div (2 ^ n) = of_bool (n = 0)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   647
  using div_mult_mod_eq [of 1 "2 ^ n"] by auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   648
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   649
lemma div_mult2_eq':
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   650
  \<open>a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   651
proof (cases \<open>m = 0 \<or> n = 0\<close>)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   652
  case True
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   653
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   654
    by auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   655
next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   656
  case False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   657
  then have \<open>m > 0\<close> \<open>n > 0\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   658
    by simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   659
  show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   660
  proof (cases \<open>of_nat m * of_nat n dvd a\<close>)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   661
    case True
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   662
    then obtain b where \<open>a = (of_nat m * of_nat n) * b\<close> ..
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   663
    then have \<open>a = of_nat m * (of_nat n * b)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   664
      by (simp add: ac_simps)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   665
    then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   666
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   667
  next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   668
    case False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   669
    define q where \<open>q = a div (of_nat m * of_nat n)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   670
    define r where \<open>r = a mod (of_nat m * of_nat n)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   671
    from \<open>m > 0\<close> \<open>n > 0\<close> \<open>\<not> of_nat m * of_nat n dvd a\<close> r_def have "division_segment r = 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   672
      using division_segment_of_nat [of "m * n"] by (simp add: division_segment_mod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   673
    with division_segment_euclidean_size [of r]
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   674
    have "of_nat (euclidean_size r) = r"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   675
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   676
    have "a mod (of_nat m * of_nat n) div (of_nat m * of_nat n) = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   677
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   678
    with \<open>m > 0\<close> \<open>n > 0\<close> r_def have "r div (of_nat m * of_nat n) = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   679
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   680
    with \<open>of_nat (euclidean_size r) = r\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   681
    have "of_nat (euclidean_size r) div (of_nat m * of_nat n) = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   682
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   683
    then have "of_nat (euclidean_size r div (m * n)) = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   684
      by (simp add: of_nat_div)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   685
    then have "of_nat (euclidean_size r div m div n) = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   686
      by (simp add: div_mult2_eq)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   687
    with \<open>of_nat (euclidean_size r) = r\<close> have "r div of_nat m div of_nat n = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   688
      by (simp add: of_nat_div)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   689
    with \<open>m > 0\<close> \<open>n > 0\<close> q_def
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   690
    have "q = (r div of_nat m + q * of_nat n * of_nat m div of_nat m) div of_nat n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   691
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   692
    moreover have \<open>a = q * (of_nat m * of_nat n) + r\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   693
      by (simp add: q_def r_def div_mult_mod_eq)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   694
    ultimately show \<open>a div (of_nat m * of_nat n) = a div of_nat m div of_nat n\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   695
      using q_def [symmetric] div_plus_div_distrib_dvd_right [of \<open>of_nat m\<close> \<open>q * (of_nat m * of_nat n)\<close> r]
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   696
      by (simp add: ac_simps)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   697
  qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   698
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   699
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   700
lemma mod_mult2_eq':
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   701
  "a mod (of_nat m * of_nat n) = of_nat m * (a div of_nat m mod of_nat n) + a mod of_nat m"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   702
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   703
  have "a div (of_nat m * of_nat n) * (of_nat m * of_nat n) + a mod (of_nat m * of_nat n) = a div of_nat m div of_nat n * of_nat n * of_nat m + (a div of_nat m mod of_nat n * of_nat m + a mod of_nat m)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   704
    by (simp add: combine_common_factor div_mult_mod_eq)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   705
  moreover have "a div of_nat m div of_nat n * of_nat n * of_nat m = of_nat n * of_nat m * (a div of_nat m div of_nat n)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   706
    by (simp add: ac_simps)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   707
  ultimately show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   708
    by (simp add: div_mult2_eq' mult_commute)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   709
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   710
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   711
lemma div_mult2_numeral_eq:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   712
  "a div numeral k div numeral l = a div numeral (k * l)" (is "?A = ?B")
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   713
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   714
  have "?A = a div of_nat (numeral k) div of_nat (numeral l)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   715
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   716
  also have "\<dots> = a div (of_nat (numeral k) * of_nat (numeral l))"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   717
    by (fact div_mult2_eq' [symmetric])
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   718
  also have "\<dots> = ?B"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   719
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   720
  finally show ?thesis .
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   721
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   722
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   723
lemma numeral_Bit0_div_2:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   724
  "numeral (num.Bit0 n) div 2 = numeral n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   725
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   726
  have "numeral (num.Bit0 n) = numeral n + numeral n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   727
    by (simp only: numeral.simps)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   728
  also have "\<dots> = numeral n * 2"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   729
    by (simp add: mult_2_right)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   730
  finally have "numeral (num.Bit0 n) div 2 = numeral n * 2 div 2"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   731
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   732
  also have "\<dots> = numeral n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   733
    by (rule nonzero_mult_div_cancel_right) simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   734
  finally show ?thesis .
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   735
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   736
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   737
lemma numeral_Bit1_div_2:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   738
  "numeral (num.Bit1 n) div 2 = numeral n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   739
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   740
  have "numeral (num.Bit1 n) = numeral n + numeral n + 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   741
    by (simp only: numeral.simps)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   742
  also have "\<dots> = numeral n * 2 + 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   743
    by (simp add: mult_2_right)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   744
  finally have "numeral (num.Bit1 n) div 2 = (numeral n * 2 + 1) div 2"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   745
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   746
  also have "\<dots> = numeral n * 2 div 2 + 1 div 2"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   747
    using dvd_triv_right by (rule div_plus_div_distrib_dvd_left)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   748
  also have "\<dots> = numeral n * 2 div 2"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   749
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   750
  also have "\<dots> = numeral n"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   751
    by (rule nonzero_mult_div_cancel_right) simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   752
  finally show ?thesis .
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   753
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   754
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   755
lemma exp_mod_exp:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   756
  \<open>2 ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   757
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   758
  have \<open>(2::nat) ^ m mod 2 ^ n = of_bool (m < n) * 2 ^ m\<close> (is \<open>?lhs = ?rhs\<close>)
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   759
    by (auto simp add: linorder_class.not_less monoid_mult_class.power_add dest!: le_Suc_ex)
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   760
  then have \<open>of_nat ?lhs = of_nat ?rhs\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   761
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   762
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   763
    by (simp add: of_nat_mod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   764
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   765
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   766
lemma mask_mod_exp:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   767
  \<open>(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - 1\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   768
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   769
  have \<open>(2 ^ n - 1) mod 2 ^ m = 2 ^ min m n - (1::nat)\<close> (is \<open>?lhs = ?rhs\<close>)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   770
  proof (cases \<open>n \<le> m\<close>)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   771
    case True
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   772
    then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   773
      by (simp add: Suc_le_lessD)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   774
  next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   775
    case False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   776
    then have \<open>m < n\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   777
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   778
    then obtain q where n: \<open>n = Suc q + m\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   779
      by (auto dest: less_imp_Suc_add)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   780
    then have \<open>min m n = m\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   781
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   782
    moreover have \<open>(2::nat) ^ m \<le> 2 * 2 ^ q * 2 ^ m\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   783
      using mult_le_mono1 [of 1 \<open>2 * 2 ^ q\<close> \<open>2 ^ m\<close>] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   784
    with n have \<open>2 ^ n - 1 = (2 ^ Suc q - 1) * 2 ^ m + (2 ^ m - (1::nat))\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   785
      by (simp add: monoid_mult_class.power_add algebra_simps)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   786
    ultimately show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   787
      by (simp only: euclidean_semiring_cancel_class.mod_mult_self3) simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   788
  qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   789
  then have \<open>of_nat ?lhs = of_nat ?rhs\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   790
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   791
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   792
    by (simp add: of_nat_mod of_nat_diff)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   793
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   794
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   795
lemma of_bool_half_eq_0 [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   796
  \<open>of_bool b div 2 = 0\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   797
  by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   798
78668
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   799
lemma of_nat_mod_double:
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   800
  \<open>of_nat n mod (2 * of_nat m) = of_nat n mod of_nat m \<or> of_nat n mod (2 * of_nat m) = of_nat n mod of_nat m + of_nat m\<close>
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   801
  by (simp add: mod_double_nat flip: of_nat_mod of_nat_add of_nat_mult of_nat_numeral)
d52934f126d4 new formulation of an auxiliary lemma
haftmann
parents: 78083
diff changeset
   802
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   803
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   804
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   805
instance nat :: linordered_euclidean_semiring
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   806
  by standard simp_all
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   807
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   808
instance int :: linordered_euclidean_semiring
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   809
  by standard (auto simp add: divide_int_def division_segment_int_def elim: contrapos_np)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   810
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   811
context linordered_euclidean_semiring
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   812
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   813
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   814
subclass semiring_parity
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   815
proof
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   816
  show \<open>a mod 2 = of_bool (\<not> 2 dvd a)\<close> for a
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   817
  proof (cases \<open>2 dvd a\<close>)
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   818
    case True
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   819
    then show ?thesis
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   820
      by (simp add: dvd_eq_mod_eq_0)
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   821
  next
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   822
    case False
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   823
    have eucl: "euclidean_size (a mod 2) = 1"
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   824
    proof (rule Orderings.order_antisym)
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   825
      show "euclidean_size (a mod 2) \<le> 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   826
        using mod_size_less [of 2 a] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   827
      show "1 \<le> euclidean_size (a mod 2)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   828
        using \<open>\<not> 2 dvd a\<close> by (simp add: Suc_le_eq dvd_eq_mod_eq_0)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   829
    qed 
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   830
    from \<open>\<not> 2 dvd a\<close> have "\<not> of_nat 2 dvd division_segment a * of_nat (euclidean_size a)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   831
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   832
    then have "\<not> of_nat 2 dvd of_nat (euclidean_size a)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   833
      by (auto simp only: dvd_mult_unit_iff' is_unit_division_segment)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   834
    then have "\<not> 2 dvd euclidean_size a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   835
      using of_nat_dvd_iff [of 2] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   836
    then have "euclidean_size a mod 2 = 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   837
      by (simp add: semidom_modulo_class.dvd_eq_mod_eq_0)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   838
    then have "of_nat (euclidean_size a mod 2) = of_nat 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   839
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   840
    then have "of_nat (euclidean_size a) mod 2 = 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   841
      by (simp add: of_nat_mod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   842
    from \<open>\<not> 2 dvd a\<close> eucl
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   843
    have "a mod 2 = 1"
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   844
      by (auto intro: division_segment_eq_iff simp add: division_segment_mod)
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   845
    with \<open>\<not> 2 dvd a\<close> show ?thesis
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   846
      by simp
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   847
  qed
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   848
  show \<open>\<not> is_unit 2\<close>
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   849
  proof
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   850
    assume \<open>is_unit 2\<close>
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   851
    then have \<open>of_nat 2 dvd of_nat 1\<close>
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   852
      by simp
79118
486a32079c60 compactified specification of type class parity
haftmann
parents: 79117
diff changeset
   853
    then have \<open>is_unit (2::nat)\<close>
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   854
      by (simp only: of_nat_dvd_iff)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   855
    then show False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   856
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   857
  qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   858
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   859
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   860
lemma even_succ_div_two [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   861
  "even a \<Longrightarrow> (a + 1) div 2 = a div 2"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   862
  by (cases "a = 0") (auto elim!: evenE dest: mult_not_zero)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   863
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   864
lemma odd_succ_div_two [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   865
  "odd a \<Longrightarrow> (a + 1) div 2 = a div 2 + 1"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   866
  by (auto elim!: oddE simp add: add.assoc)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   867
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   868
lemma even_two_times_div_two:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   869
  "even a \<Longrightarrow> 2 * (a div 2) = a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   870
  by (fact dvd_mult_div_cancel)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   871
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   872
lemma odd_two_times_div_two_succ [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   873
  "odd a \<Longrightarrow> 2 * (a div 2) + 1 = a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   874
  using mult_div_mod_eq [of 2 a]
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   875
  by (simp add: even_iff_mod_2_eq_zero)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   876
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   877
lemma coprime_left_2_iff_odd [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   878
  "coprime 2 a \<longleftrightarrow> odd a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   879
proof
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   880
  assume "odd a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   881
  show "coprime 2 a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   882
  proof (rule coprimeI)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   883
    fix b
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   884
    assume "b dvd 2" "b dvd a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   885
    then have "b dvd a mod 2"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   886
      by (auto intro: dvd_mod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   887
    with \<open>odd a\<close> show "is_unit b"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   888
      by (simp add: mod_2_eq_odd)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   889
  qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   890
next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   891
  assume "coprime 2 a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   892
  show "odd a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   893
  proof (rule notI)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   894
    assume "even a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   895
    then obtain b where "a = 2 * b" ..
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   896
    with \<open>coprime 2 a\<close> have "coprime 2 (2 * b)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   897
      by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   898
    moreover have "\<not> coprime 2 (2 * b)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   899
      by (rule not_coprimeI [of 2]) simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   900
    ultimately show False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   901
      by blast
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   902
  qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   903
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   904
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   905
lemma coprime_right_2_iff_odd [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   906
  "coprime a 2 \<longleftrightarrow> odd a"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   907
  using coprime_left_2_iff_odd [of a] by (simp add: ac_simps)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   908
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   909
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   910
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   911
lemma minus_1_mod_2_eq [simp]:
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   912
  \<open>- 1 mod 2 = (1::int)\<close>
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   913
  by (simp add: mod_2_eq_odd)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   914
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   915
lemma minus_1_div_2_eq [simp]:
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   916
  "- 1 div 2 = - (1::int)"
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   917
proof -
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   918
  from div_mult_mod_eq [of "- 1 :: int" 2]
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   919
  have "- 1 div 2 * 2 = - 1 * (2 :: int)"
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   920
    using add_implies_diff by fastforce
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   921
  then show ?thesis
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   922
    using mult_right_cancel [of 2 "- 1 div 2" "- (1 :: int)"] by simp
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   923
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   924
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   925
context linordered_euclidean_semiring
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   926
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   927
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   928
lemma even_mask_div_iff':
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   929
  \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> m \<le> n\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   930
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   931
  have \<open>even ((2 ^ m - 1) div 2 ^ n) \<longleftrightarrow> even (of_nat ((2 ^ m - Suc 0) div 2 ^ n))\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   932
    by (simp only: of_nat_div) (simp add: of_nat_diff)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   933
  also have \<open>\<dots> \<longleftrightarrow> even ((2 ^ m - Suc 0) div 2 ^ n)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   934
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   935
  also have \<open>\<dots> \<longleftrightarrow> m \<le> n\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   936
  proof (cases \<open>m \<le> n\<close>)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   937
    case True
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   938
    then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   939
      by (simp add: Suc_le_lessD)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   940
  next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   941
    case False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   942
    then obtain r where r: \<open>m = n + Suc r\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   943
      using less_imp_Suc_add by fastforce
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   944
    from r have \<open>{q. q < m} \<inter> {q. 2 ^ n dvd (2::nat) ^ q} = {q. n \<le> q \<and> q < m}\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   945
      by (auto simp add: dvd_power_iff_le)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   946
    moreover from r have \<open>{q. q < m} \<inter> {q. \<not> 2 ^ n dvd (2::nat) ^ q} = {q. q < n}\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   947
      by (auto simp add: dvd_power_iff_le)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   948
    moreover from False have \<open>{q. n \<le> q \<and> q < m \<and> q \<le> n} = {n}\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   949
      by auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   950
    then have \<open>odd ((\<Sum>a\<in>{q. n \<le> q \<and> q < m}. 2 ^ a div (2::nat) ^ n) + sum ((^) 2) {q. q < n} div 2 ^ n)\<close>
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   951
      by (simp_all add: euclidean_semiring_cancel_class.power_diff_power_eq semiring_parity_class.even_sum_iff
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   952
        linorder_class.not_less mask_eq_sum_exp_nat [symmetric])
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   953
    ultimately have \<open>odd (sum ((^) (2::nat)) {q. q < m} div 2 ^ n)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   954
      by (subst euclidean_semiring_cancel_class.sum_div_partition) simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   955
    with False show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   956
      by (simp add: mask_eq_sum_exp_nat)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   957
  qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   958
  finally show ?thesis .
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   959
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   960
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   961
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   962
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   963
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   964
subsection \<open>Generic symbolic computations\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   965
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   966
text \<open>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   967
  The following type class contains everything necessary to formulate
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   968
  a division algorithm in ring structures with numerals, restricted
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   969
  to its positive segments.
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   970
\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   971
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
   972
class linordered_euclidean_semiring_division = linordered_euclidean_semiring +
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   973
  fixes divmod :: \<open>num \<Rightarrow> num \<Rightarrow> 'a \<times> 'a\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   974
    and divmod_step :: \<open>'a \<Rightarrow> 'a \<times> 'a \<Rightarrow> 'a \<times> 'a\<close> \<comment> \<open>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   975
      These are conceptually definitions but force generated code
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   976
      to be monomorphic wrt. particular instances of this class which
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   977
      yields a significant speedup.\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   978
  assumes divmod_def: \<open>divmod m n = (numeral m div numeral n, numeral m mod numeral n)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   979
    and divmod_step_def [simp]: \<open>divmod_step l (q, r) =
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   980
      (if euclidean_size l \<le> euclidean_size r then (2 * q + 1, r - l)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   981
       else (2 * q, r))\<close> \<comment> \<open>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   982
         This is a formulation of one step (referring to one digit position)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   983
         in school-method division: compare the dividend at the current
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   984
         digit position with the remainder from previous division steps
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   985
         and evaluate accordingly.\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   986
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   987
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   988
lemma fst_divmod:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   989
  \<open>fst (divmod m n) = numeral m div numeral n\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   990
  by (simp add: divmod_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   991
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   992
lemma snd_divmod:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   993
  \<open>snd (divmod m n) = numeral m mod numeral n\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   994
  by (simp add: divmod_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   995
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   996
text \<open>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   997
  Following a formulation of school-method division.
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   998
  If the divisor is smaller than the dividend, terminate.
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
   999
  If not, shift the dividend to the right until termination
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1000
  occurs and then reiterate single division steps in the
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1001
  opposite direction.
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1002
\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1003
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1004
lemma divmod_divmod_step:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1005
  \<open>divmod m n = (if m < n then (0, numeral m)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1006
    else divmod_step (numeral n) (divmod m (Num.Bit0 n)))\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1007
proof (cases \<open>m < n\<close>)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1008
  case True
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1009
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1010
    by (simp add: prod_eq_iff fst_divmod snd_divmod flip: of_nat_numeral of_nat_div of_nat_mod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1011
next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1012
  case False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1013
  define r s t where \<open>r = (numeral m :: nat)\<close> \<open>s = (numeral n :: nat)\<close> \<open>t = 2 * s\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1014
  then have *: \<open>numeral m = of_nat r\<close> \<open>numeral n = of_nat s\<close> \<open>numeral (num.Bit0 n) = of_nat t\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1015
    and \<open>\<not> s \<le> r mod s\<close>
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1016
    by (simp_all add: linorder_class.not_le)
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1017
  have t: \<open>2 * (r div t) = r div s - r div s mod 2\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1018
    \<open>r mod t = s * (r div s mod 2) + r mod s\<close>
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 76387
diff changeset
  1019
    by (simp add: Rings.minus_mod_eq_mult_div Groups.mult.commute [of 2] Euclidean_Rings.div_mult2_eq \<open>t = 2 * s\<close>)
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1020
      (use mod_mult2_eq [of r s 2] in \<open>simp add: ac_simps \<open>t = 2 * s\<close>\<close>)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1021
  have rs: \<open>r div s mod 2 = 0 \<or> r div s mod 2 = Suc 0\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1022
    by auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1023
  from \<open>\<not> s \<le> r mod s\<close> have \<open>s \<le> r mod t \<Longrightarrow>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1024
     r div s = Suc (2 * (r div t)) \<and>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1025
     r mod s = r mod t - s\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1026
    using rs
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1027
    by (auto simp add: t)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1028
  moreover have \<open>r mod t < s \<Longrightarrow>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1029
     r div s = 2 * (r div t) \<and>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1030
     r mod s = r mod t\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1031
    using rs
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1032
    by (auto simp add: t)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1033
  ultimately show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1034
    by (simp add: divmod_def prod_eq_iff split_def Let_def
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1035
        not_less mod_eq_0_iff_dvd Rings.mod_eq_0_iff_dvd False not_le *)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1036
    (simp add: flip: of_nat_numeral of_nat_mult add.commute [of 1] of_nat_div of_nat_mod of_nat_Suc of_nat_diff)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1037
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1038
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1039
text \<open>The division rewrite proper -- first, trivial results involving \<open>1\<close>\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1040
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1041
lemma divmod_trivial [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1042
  "divmod m Num.One = (numeral m, 0)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1043
  "divmod num.One (num.Bit0 n) = (0, Numeral1)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1044
  "divmod num.One (num.Bit1 n) = (0, Numeral1)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1045
  using divmod_divmod_step [of "Num.One"] by (simp_all add: divmod_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1046
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1047
text \<open>Division by an even number is a right-shift\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1048
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1049
lemma divmod_cancel [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1050
  \<open>divmod (Num.Bit0 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r))\<close> (is ?P)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1051
  \<open>divmod (Num.Bit1 m) (Num.Bit0 n) = (case divmod m n of (q, r) \<Rightarrow> (q, 2 * r + 1))\<close> (is ?Q)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1052
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1053
  define r s where \<open>r = (numeral m :: nat)\<close> \<open>s = (numeral n :: nat)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1054
  then have *: \<open>numeral m = of_nat r\<close> \<open>numeral n = of_nat s\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1055
    \<open>numeral (num.Bit0 m) = of_nat (2 * r)\<close> \<open>numeral (num.Bit0 n) = of_nat (2 * s)\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1056
    \<open>numeral (num.Bit1 m) = of_nat (Suc (2 * r))\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1057
    by simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1058
  have **: \<open>Suc (2 * r) div 2 = r\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1059
    by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1060
  show ?P and ?Q
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1061
    by (simp_all add: divmod_def *)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1062
      (simp_all flip: of_nat_numeral of_nat_div of_nat_mod of_nat_mult add.commute [of 1] of_nat_Suc
77061
5de3772609ea generalized theory name: euclidean division denotes one particular division definition on integers
haftmann
parents: 76387
diff changeset
  1063
       add: Euclidean_Rings.mod_mult_mult1 div_mult2_eq [of _ 2] mod_mult2_eq [of _ 2] **)
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1064
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1065
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1066
text \<open>The really hard work\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1067
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1068
lemma divmod_steps [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1069
  "divmod (num.Bit0 m) (num.Bit1 n) =
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1070
      (if m \<le> n then (0, numeral (num.Bit0 m))
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1071
       else divmod_step (numeral (num.Bit1 n))
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1072
             (divmod (num.Bit0 m)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1073
               (num.Bit0 (num.Bit1 n))))"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1074
  "divmod (num.Bit1 m) (num.Bit1 n) =
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1075
      (if m < n then (0, numeral (num.Bit1 m))
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1076
       else divmod_step (numeral (num.Bit1 n))
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1077
             (divmod (num.Bit1 m)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1078
               (num.Bit0 (num.Bit1 n))))"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1079
  by (simp_all add: divmod_divmod_step)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1080
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1081
lemmas divmod_algorithm_code = divmod_trivial divmod_cancel divmod_steps
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1082
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1083
text \<open>Special case: divisibility\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1084
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1085
definition divides_aux :: "'a \<times> 'a \<Rightarrow> bool"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1086
where
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1087
  "divides_aux qr \<longleftrightarrow> snd qr = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1088
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1089
lemma divides_aux_eq [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1090
  "divides_aux (q, r) \<longleftrightarrow> r = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1091
  by (simp add: divides_aux_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1092
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1093
lemma dvd_numeral_simp [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1094
  "numeral m dvd numeral n \<longleftrightarrow> divides_aux (divmod n m)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1095
  by (simp add: divmod_def mod_eq_0_iff_dvd)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1096
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1097
text \<open>Generic computation of quotient and remainder\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1098
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1099
lemma numeral_div_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1100
  "numeral k div numeral l = fst (divmod k l)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1101
  by (simp add: fst_divmod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1102
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1103
lemma numeral_mod_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1104
  "numeral k mod numeral l = snd (divmod k l)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1105
  by (simp add: snd_divmod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1106
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1107
lemma one_div_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1108
  "1 div numeral n = fst (divmod num.One n)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1109
  by (simp add: fst_divmod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1110
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1111
lemma one_mod_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1112
  "1 mod numeral n = snd (divmod num.One n)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1113
  by (simp add: snd_divmod)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1114
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1115
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1116
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1117
instantiation nat :: linordered_euclidean_semiring_division
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1118
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1119
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1120
definition divmod_nat :: "num \<Rightarrow> num \<Rightarrow> nat \<times> nat"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1121
where
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1122
  divmod'_nat_def: "divmod_nat m n = (numeral m div numeral n, numeral m mod numeral n)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1123
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1124
definition divmod_step_nat :: "nat \<Rightarrow> nat \<times> nat \<Rightarrow> nat \<times> nat"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1125
where
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1126
  "divmod_step_nat l qr = (let (q, r) = qr
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1127
    in if r \<ge> l then (2 * q + 1, r - l)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1128
    else (2 * q, r))"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1129
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1130
instance
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1131
  by standard (simp_all add: divmod'_nat_def divmod_step_nat_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1132
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1133
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1134
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1135
declare divmod_algorithm_code [where ?'a = nat, code]
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1136
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1137
lemma Suc_0_div_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1138
  \<open>Suc 0 div numeral Num.One = 1\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1139
  \<open>Suc 0 div numeral (Num.Bit0 n) = 0\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1140
  \<open>Suc 0 div numeral (Num.Bit1 n) = 0\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1141
  by simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1142
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1143
lemma Suc_0_mod_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1144
  \<open>Suc 0 mod numeral Num.One = 0\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1145
  \<open>Suc 0 mod numeral (Num.Bit0 n) = 1\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1146
  \<open>Suc 0 mod numeral (Num.Bit1 n) = 1\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1147
  by simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1148
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1149
instantiation int :: linordered_euclidean_semiring_division
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1150
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1151
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1152
definition divmod_int :: "num \<Rightarrow> num \<Rightarrow> int \<times> int"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1153
where
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1154
  "divmod_int m n = (numeral m div numeral n, numeral m mod numeral n)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1155
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1156
definition divmod_step_int :: "int \<Rightarrow> int \<times> int \<Rightarrow> int \<times> int"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1157
where
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1158
  "divmod_step_int l qr = (let (q, r) = qr
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1159
    in if \<bar>l\<bar> \<le> \<bar>r\<bar> then (2 * q + 1, r - l)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1160
    else (2 * q, r))"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1161
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1162
instance
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1163
  by standard (auto simp add: divmod_int_def divmod_step_int_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1164
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1165
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1166
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1167
declare divmod_algorithm_code [where ?'a = int, code]
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1168
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1169
context
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1170
begin
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1171
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1172
qualified definition adjust_div :: "int \<times> int \<Rightarrow> int"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1173
where
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1174
  "adjust_div qr = (let (q, r) = qr in q + of_bool (r \<noteq> 0))"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1175
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1176
qualified lemma adjust_div_eq [simp, code]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1177
  "adjust_div (q, r) = q + of_bool (r \<noteq> 0)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1178
  by (simp add: adjust_div_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1179
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1180
qualified definition adjust_mod :: "num \<Rightarrow> int \<Rightarrow> int"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1181
where
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1182
  [simp]: "adjust_mod l r = (if r = 0 then 0 else numeral l - r)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1183
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1184
lemma minus_numeral_div_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1185
  "- numeral m div numeral n = - (adjust_div (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1186
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1187
  have "int (fst (divmod m n)) = fst (divmod m n)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1188
    by (simp only: fst_divmod divide_int_def) auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1189
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1190
    by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1191
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1192
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1193
lemma minus_numeral_mod_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1194
  "- numeral m mod numeral n = adjust_mod n (snd (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1195
proof (cases "snd (divmod m n) = (0::int)")
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1196
  case True
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1197
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1198
    by (simp add: mod_eq_0_iff_dvd divides_aux_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1199
next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1200
  case False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1201
  then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1202
    by (simp only: snd_divmod modulo_int_def) auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1203
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1204
    by (simp add: divides_aux_def adjust_div_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1205
      (simp add: divides_aux_def modulo_int_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1206
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1207
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1208
lemma numeral_div_minus_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1209
  "numeral m div - numeral n = - (adjust_div (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1210
proof -
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1211
  have "int (fst (divmod m n)) = fst (divmod m n)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1212
    by (simp only: fst_divmod divide_int_def) auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1213
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1214
    by (auto simp add: split_def Let_def adjust_div_def divides_aux_def divide_int_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1215
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1216
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1217
lemma numeral_mod_minus_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1218
  "numeral m mod - numeral n = - adjust_mod n (snd (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1219
proof (cases "snd (divmod m n) = (0::int)")
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1220
  case True
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1221
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1222
    by (simp add: mod_eq_0_iff_dvd divides_aux_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1223
next
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1224
  case False
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1225
  then have "int (snd (divmod m n)) = snd (divmod m n)" if "snd (divmod m n) \<noteq> (0::int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1226
    by (simp only: snd_divmod modulo_int_def) auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1227
  then show ?thesis
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1228
    by (simp add: divides_aux_def adjust_div_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1229
      (simp add: divides_aux_def modulo_int_def)
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1230
qed
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1231
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1232
lemma minus_one_div_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1233
  "- 1 div numeral n = - (adjust_div (divmod Num.One n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1234
  using minus_numeral_div_numeral [of Num.One n] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1235
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1236
lemma minus_one_mod_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1237
  "- 1 mod numeral n = adjust_mod n (snd (divmod Num.One n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1238
  using minus_numeral_mod_numeral [of Num.One n] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1239
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1240
lemma one_div_minus_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1241
  "1 div - numeral n = - (adjust_div (divmod Num.One n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1242
  using numeral_div_minus_numeral [of Num.One n] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1243
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1244
lemma one_mod_minus_numeral [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1245
  "1 mod - numeral n = - adjust_mod n (snd (divmod Num.One n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1246
  using numeral_mod_minus_numeral [of Num.One n] by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1247
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1248
lemma [code]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1249
  fixes k :: int
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1250
  shows
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1251
    "k div 0 = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1252
    "k mod 0 = k"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1253
    "0 div k = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1254
    "0 mod k = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1255
    "k div Int.Pos Num.One = k"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1256
    "k mod Int.Pos Num.One = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1257
    "k div Int.Neg Num.One = - k"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1258
    "k mod Int.Neg Num.One = 0"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1259
    "Int.Pos m div Int.Pos n = (fst (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1260
    "Int.Pos m mod Int.Pos n = (snd (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1261
    "Int.Neg m div Int.Pos n = - (adjust_div (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1262
    "Int.Neg m mod Int.Pos n = adjust_mod n (snd (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1263
    "Int.Pos m div Int.Neg n = - (adjust_div (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1264
    "Int.Pos m mod Int.Neg n = - adjust_mod n (snd (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1265
    "Int.Neg m div Int.Neg n = (fst (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1266
    "Int.Neg m mod Int.Neg n = - (snd (divmod m n) :: int)"
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1267
  by simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1268
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1269
end
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1270
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1271
lemma divmod_BitM_2_eq [simp]:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1272
  \<open>divmod (Num.BitM m) (Num.Bit0 Num.One) = (numeral m - 1, (1 :: int))\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1273
  by (cases m) simp_all
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1274
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1275
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1276
subsubsection \<open>Computation by simplification\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1277
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1278
lemma euclidean_size_nat_less_eq_iff:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1279
  \<open>euclidean_size m \<le> euclidean_size n \<longleftrightarrow> m \<le> n\<close> for m n :: nat
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1280
  by simp
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1281
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1282
lemma euclidean_size_int_less_eq_iff:
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1283
  \<open>euclidean_size k \<le> euclidean_size l \<longleftrightarrow> \<bar>k\<bar> \<le> \<bar>l\<bar>\<close> for k l :: int
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1284
  by auto
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1285
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1286
simproc_setup numeral_divmod
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1287
  ("0 div 0 :: 'a :: linordered_euclidean_semiring_division" | "0 mod 0 :: 'a :: linordered_euclidean_semiring_division" |
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1288
   "0 div 1 :: 'a :: linordered_euclidean_semiring_division" | "0 mod 1 :: 'a :: linordered_euclidean_semiring_division" |
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1289
   "0 div - 1 :: int" | "0 mod - 1 :: int" |
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1290
   "0 div numeral b :: 'a :: linordered_euclidean_semiring_division" | "0 mod numeral b :: 'a :: linordered_euclidean_semiring_division" |
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1291
   "0 div - numeral b :: int" | "0 mod - numeral b :: int" |
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1292
   "1 div 0 :: 'a :: linordered_euclidean_semiring_division" | "1 mod 0 :: 'a :: linordered_euclidean_semiring_division" |
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1293
   "1 div 1 :: 'a :: linordered_euclidean_semiring_division" | "1 mod 1 :: 'a :: linordered_euclidean_semiring_division" |
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1294
   "1 div - 1 :: int" | "1 mod - 1 :: int" |
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1295
   "1 div numeral b :: 'a :: linordered_euclidean_semiring_division" | "1 mod numeral b :: 'a :: linordered_euclidean_semiring_division" |
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1296
   "1 div - numeral b :: int" |"1 mod - numeral b :: int" |
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1297
   "- 1 div 0 :: int" | "- 1 mod 0 :: int" | "- 1 div 1 :: int" | "- 1 mod 1 :: int" |
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1298
   "- 1 div - 1 :: int" | "- 1 mod - 1 :: int" | "- 1 div numeral b :: int" | "- 1 mod numeral b :: int" |
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1299
   "- 1 div - numeral b :: int" | "- 1 mod - numeral b :: int" |
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1300
   "numeral a div 0 :: 'a :: linordered_euclidean_semiring_division" | "numeral a mod 0 :: 'a :: linordered_euclidean_semiring_division" |
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1301
   "numeral a div 1 :: 'a :: linordered_euclidean_semiring_division" | "numeral a mod 1 :: 'a :: linordered_euclidean_semiring_division" |
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1302
   "numeral a div - 1 :: int" | "numeral a mod - 1 :: int" |
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1303
   "numeral a div numeral b :: 'a :: linordered_euclidean_semiring_division" | "numeral a mod numeral b :: 'a :: linordered_euclidean_semiring_division" |
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1304
   "numeral a div - numeral b :: int" | "numeral a mod - numeral b :: int" |
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1305
   "- numeral a div 0 :: int" | "- numeral a mod 0 :: int" |
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1306
   "- numeral a div 1 :: int" | "- numeral a mod 1 :: int" |
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1307
   "- numeral a div - 1 :: int" | "- numeral a mod - 1 :: int" |
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1308
   "- numeral a div numeral b :: int" | "- numeral a mod numeral b :: int" |
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1309
   "- numeral a div - numeral b :: int" | "- numeral a mod - numeral b :: int") = \<open>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1310
  let
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1311
    val if_cong = the (Code.get_case_cong \<^theory> \<^const_name>\<open>If\<close>);
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1312
    fun successful_rewrite ctxt ct =
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1313
      let
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1314
        val thm = Simplifier.rewrite ctxt ct
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1315
      in if Thm.is_reflexive thm then NONE else SOME thm end;
78082
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1316
    val simps = @{thms div_0 mod_0 div_by_0 mod_by_0 div_by_1 mod_by_1
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1317
      one_div_numeral one_mod_numeral minus_one_div_numeral minus_one_mod_numeral
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1318
      one_div_minus_numeral one_mod_minus_numeral
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1319
      numeral_div_numeral numeral_mod_numeral minus_numeral_div_numeral minus_numeral_mod_numeral
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1320
      numeral_div_minus_numeral numeral_mod_minus_numeral
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1321
      div_minus_minus mod_minus_minus Parity.adjust_div_eq of_bool_eq one_neq_zero
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1322
      numeral_neq_zero neg_equal_0_iff_equal arith_simps arith_special divmod_trivial
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1323
      divmod_cancel divmod_steps divmod_step_def fst_conv snd_conv numeral_One
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1324
      case_prod_beta rel_simps Parity.adjust_mod_def div_minus1_right mod_minus1_right
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1325
      minus_minus numeral_times_numeral mult_zero_right mult_1_right
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1326
      euclidean_size_nat_less_eq_iff euclidean_size_int_less_eq_iff diff_nat_numeral nat_numeral}
a51d2e96203e omit pointless morphism in global theory;
wenzelm
parents: 77061
diff changeset
  1327
      @ [@{lemma "0 = 0 \<longleftrightarrow> True" by simp}];
78083
3357bc875b11 prefer static simpset;
wenzelm
parents: 78082
diff changeset
  1328
    val simpset =
3357bc875b11 prefer static simpset;
wenzelm
parents: 78082
diff changeset
  1329
      HOL_ss |> Simplifier.simpset_map \<^context>
3357bc875b11 prefer static simpset;
wenzelm
parents: 78082
diff changeset
  1330
        (Simplifier.add_cong if_cong #> fold Simplifier.add_simp simps);
3357bc875b11 prefer static simpset;
wenzelm
parents: 78082
diff changeset
  1331
  in K (fn ctxt => successful_rewrite (Simplifier.put_simpset simpset ctxt)) end
76387
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1332
\<close> \<comment> \<open>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1333
  There is space for improvement here: the calculation itself
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1334
  could be carried out outside the logic, and a generic simproc
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1335
  (simplifier setup) for generic calculation would be helpful.
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1336
\<close>
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1337
8cb141384682 restructured
haftmann
parents: 75937
diff changeset
  1338
75937
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1339
subsection \<open>Computing congruences modulo \<open>2 ^ q\<close>\<close>
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1340
78937
5e6b195eee83 slightly less technical formulation of very specific type class
haftmann
parents: 78668
diff changeset
  1341
context linordered_euclidean_semiring_division
75937
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1342
begin
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1343
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1344
lemma cong_exp_iff_simps:
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1345
  "numeral n mod numeral Num.One = 0
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1346
    \<longleftrightarrow> True"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1347
  "numeral (Num.Bit0 n) mod numeral (Num.Bit0 q) = 0
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1348
    \<longleftrightarrow> numeral n mod numeral q = 0"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1349
  "numeral (Num.Bit1 n) mod numeral (Num.Bit0 q) = 0
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1350
    \<longleftrightarrow> False"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1351
  "numeral m mod numeral Num.One = (numeral n mod numeral Num.One)
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1352
    \<longleftrightarrow> True"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1353
  "numeral Num.One mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1354
    \<longleftrightarrow> True"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1355
  "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1356
    \<longleftrightarrow> False"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1357
  "numeral Num.One mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1358
    \<longleftrightarrow> (numeral n mod numeral q) = 0"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1359
  "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1360
    \<longleftrightarrow> False"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1361
  "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1362
    \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1363
  "numeral (Num.Bit0 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1364
    \<longleftrightarrow> False"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1365
  "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral Num.One mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1366
    \<longleftrightarrow> (numeral m mod numeral q) = 0"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1367
  "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit0 n) mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1368
    \<longleftrightarrow> False"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1369
  "numeral (Num.Bit1 m) mod numeral (Num.Bit0 q) = (numeral (Num.Bit1 n) mod numeral (Num.Bit0 q))
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1370
    \<longleftrightarrow> numeral m mod numeral q = (numeral n mod numeral q)"
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1371
  by (auto simp add: case_prod_beta dest: arg_cong [of _ _ even])
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1372
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1373
end
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1374
02b18f59f903 streamlined
haftmann
parents: 74592
diff changeset
  1375
71853
30d92e668b52 tuned module name space for generated code
haftmann
parents: 71837
diff changeset
  1376
code_identifier
30d92e668b52 tuned module name space for generated code
haftmann
parents: 71837
diff changeset
  1377
  code_module Parity \<rightharpoonup> (SML) Arith and (OCaml) Arith and (Haskell) Arith
30d92e668b52 tuned module name space for generated code
haftmann
parents: 71837
diff changeset
  1378
74592
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74101
diff changeset
  1379
lemmas even_of_nat = even_of_nat_iff
3c587b7c3d5c more generic bit/word lemmas for distribution
haftmann
parents: 74101
diff changeset
  1380
67816
2249b27ab1dd abstract algebraic bit operations
haftmann
parents: 67371
diff changeset
  1381
end