src/HOL/Library/Discrete_Functions.thy
author nipkow
Tue, 17 Jun 2025 14:11:40 +0200
changeset 82733 8b537e1af2ec
parent 82664 e9f3b94eb6a0
permissions -rw-r--r--
reinstated intersection of lists as inter_list_set
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
59349
3bde948f439c tuned -- more Sidekick-friendly layout;
wenzelm
parents: 58881
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     2
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59349
diff changeset
     3
section \<open>Common discrete functions\<close>
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     4
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
     5
theory Discrete_Functions
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
     6
imports Complex_Main
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     7
begin
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     8
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59349
diff changeset
     9
subsection \<open>Discrete logarithm\<close>
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    10
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    11
fun floor_log :: "nat \<Rightarrow> nat"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    12
  where [simp del]: "floor_log n = (if n < 2 then 0 else Suc (floor_log (n div 2)))"
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60500
diff changeset
    13
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    14
lemma floor_log_induct [consumes 1, case_names one double]:
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    15
  fixes n :: nat
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    16
  assumes "n > 0"
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    17
  assumes one: "P 1"
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    18
  assumes double: "\<And>n. n \<ge> 2 \<Longrightarrow> P (n div 2) \<Longrightarrow> P n"
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    19
  shows "P n"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    20
using \<open>n > 0\<close> proof (induct n rule: floor_log.induct)
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    21
  fix n
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    22
  assume "\<not> n < 2 \<Longrightarrow>
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    23
          0 < n div 2 \<Longrightarrow> P (n div 2)"
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    24
  then have *: "n \<ge> 2 \<Longrightarrow> P (n div 2)" by simp
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    25
  assume "n > 0"
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    26
  show "P n"
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    27
  proof (cases "n = 1")
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63516
diff changeset
    28
    case True
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63516
diff changeset
    29
    with one show ?thesis by simp
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    30
  next
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63516
diff changeset
    31
    case False
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63516
diff changeset
    32
    with \<open>n > 0\<close> have "n \<ge> 2" by auto
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63516
diff changeset
    33
    with * have "P (n div 2)" .
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63516
diff changeset
    34
    with \<open>n \<ge> 2\<close> show ?thesis by (rule double)
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    35
  qed
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    36
qed
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 80621
diff changeset
    37
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    38
lemma floor_log_zero [simp]: "floor_log 0 = 0"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    39
  by (simp add: floor_log.simps)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    40
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    41
lemma floor_log_one [simp]: "floor_log 1 = 0"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    42
  by (simp add: floor_log.simps)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    43
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    44
lemma floor_log_Suc_zero [simp]: "floor_log (Suc 0) = 0"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    45
  using floor_log_one by simp
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    46
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    47
lemma floor_log_rec: "n \<ge> 2 \<Longrightarrow> floor_log n = Suc (floor_log (n div 2))"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    48
  by (simp add: floor_log.simps)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    49
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    50
lemma floor_log_twice [simp]: "n \<noteq> 0 \<Longrightarrow> floor_log (2 * n) = Suc (floor_log n)"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    51
  by (simp add: floor_log_rec)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    52
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    53
lemma floor_log_half [simp]: "floor_log (n div 2) = floor_log n - 1"
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    54
proof (cases "n < 2")
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    55
  case True
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    56
  then have "n = 0 \<or> n = 1" by arith
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    57
  then show ?thesis by (auto simp del: One_nat_def)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    58
next
59349
3bde948f439c tuned -- more Sidekick-friendly layout;
wenzelm
parents: 58881
diff changeset
    59
  case False
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    60
  then show ?thesis by (simp add: floor_log_rec)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    61
qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    62
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    63
lemma floor_log_power [simp]: "floor_log (2 ^ n) = n"
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    64
  by (induct n) simp_all
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    65
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    66
lemma floor_log_mono: "mono floor_log"
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    67
proof
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    68
  fix m n :: nat
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    69
  assume "m \<le> n"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    70
  then show "floor_log m \<le> floor_log n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    71
  proof (induct m arbitrary: n rule: floor_log.induct)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    72
    case (1 m)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    73
    then have mn2: "m div 2 \<le> n div 2" by arith
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    74
    show "floor_log m \<le> floor_log n"
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    75
    proof (cases "m \<ge> 2")
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    76
      case False
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    77
      then have "m = 0 \<or> m = 1" by arith
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    78
      then show ?thesis by (auto simp del: One_nat_def)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    79
    next
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    80
      case True then have "\<not> m < 2" by simp
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    81
      with mn2 have "n \<ge> 2" by arith
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    82
      from True have m2_0: "m div 2 \<noteq> 0" by arith
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    83
      with mn2 have n2_0: "n div 2 \<noteq> 0" by arith
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    84
      from \<open>\<not> m < 2\<close> "1.hyps" mn2 have "floor_log (m div 2) \<le> floor_log (n div 2)" by blast
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    85
      with m2_0 n2_0 have "floor_log (2 * (m div 2)) \<le> floor_log (2 * (n div 2))" by simp
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    86
      with m2_0 n2_0 \<open>m \<ge> 2\<close> \<open>n \<ge> 2\<close> show ?thesis by (simp only: floor_log_rec [of m] floor_log_rec [of n]) simp
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    87
    qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    88
  qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    89
qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    90
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    91
lemma floor_log_exp2_le:
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    92
  assumes "n > 0"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    93
  shows "2 ^ floor_log n \<le> n"
63516
76492eaf3dc1 tuned proofs;
wenzelm
parents: 62390
diff changeset
    94
  using assms
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
    95
proof (induct n rule: floor_log_induct)
63516
76492eaf3dc1 tuned proofs;
wenzelm
parents: 62390
diff changeset
    96
  case one
76492eaf3dc1 tuned proofs;
wenzelm
parents: 62390
diff changeset
    97
  then show ?case by simp
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
    98
next
63516
76492eaf3dc1 tuned proofs;
wenzelm
parents: 62390
diff changeset
    99
  case (double n)
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   100
  with floor_log_mono have "floor_log n \<ge> Suc 0"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   101
    by (simp add: floor_log.simps)
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   102
  assume "2 ^ floor_log (n div 2) \<le> n div 2"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   103
  with \<open>n \<ge> 2\<close> have "2 ^ (floor_log n - Suc 0) \<le> n div 2" by simp
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   104
  then have "2 ^ (floor_log n - Suc 0) * 2 ^ 1 \<le> n div 2 * 2" by simp
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   105
  with \<open>floor_log n \<ge> Suc 0\<close> have "2 ^ floor_log n \<le> n div 2 * 2"
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
   106
    unfolding power_add [symmetric] by simp
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
   107
  also have "n div 2 * 2 \<le> n" by (cases "even n") simp_all
63516
76492eaf3dc1 tuned proofs;
wenzelm
parents: 62390
diff changeset
   108
  finally show ?case .
61831
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
   109
qed
c43f87119d80 modernized
haftmann
parents: 61115
diff changeset
   110
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   111
lemma floor_log_exp2_gt: "2 * 2 ^ floor_log n > n"
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   112
proof (cases "n > 0")
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   113
  case True
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   114
  thus ?thesis
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   115
  proof (induct n rule: floor_log_induct)
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   116
    case (double n)
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   117
    thus ?case
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   118
      by (cases "even n") (auto elim!: evenE oddE simp: field_simps floor_log.simps)
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   119
  qed simp_all
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   120
qed simp_all
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   121
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   122
lemma floor_log_exp2_ge: "2 * 2 ^ floor_log n \<ge> n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   123
  using floor_log_exp2_gt[of n] by simp
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   124
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   125
lemma floor_log_le_iff: "m \<le> n \<Longrightarrow> floor_log m \<le> floor_log n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   126
  by (rule monoD [OF floor_log_mono])
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   127
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   128
lemma floor_log_eqI:
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   129
  assumes "n > 0" "2^k \<le> n" "n < 2 * 2^k"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   130
  shows   "floor_log n = k"
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   131
proof (rule antisym)
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   132
  from \<open>n > 0\<close> have "2 ^ floor_log n \<le> n" by (rule floor_log_exp2_le)
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   133
  also have "\<dots> < 2 ^ Suc k" using assms by simp
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   134
  finally have "floor_log n < Suc k" by (subst (asm) power_strict_increasing_iff) simp_all
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   135
  thus "floor_log n \<le> k" by simp
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   136
next
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   137
  have "2^k \<le> n" by fact
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   138
  also have "\<dots> < 2^(Suc (floor_log n))" by (simp add: floor_log_exp2_gt)
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   139
  finally have "k < Suc (floor_log n)" by (subst (asm) power_strict_increasing_iff) simp_all
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   140
  thus "k \<le> floor_log n" by simp
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   141
qed
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   142
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   143
lemma floor_log_altdef: "floor_log n = (if n = 0 then 0 else nat \<lfloor>log 2 (real_of_nat n)\<rfloor>)"
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   144
proof (cases "n = 0")
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   145
  case False
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   146
  have "\<lfloor>log 2 (real_of_nat n)\<rfloor> = int (floor_log n)"
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   147
  proof (rule floor_unique)
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   148
    from False have "2 powr (real (floor_log n)) \<le> real n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   149
      by (simp add: powr_realpow floor_log_exp2_le)
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   150
    hence "log 2 (2 powr (real (floor_log n))) \<le> log 2 (real n)"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   151
      using False by (subst log_le_cancel_iff) simp_all
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   152
    also have "log 2 (2 powr (real (floor_log n))) = real (floor_log n)" by simp
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   153
    finally show "real_of_int (int (floor_log n)) \<le> log 2 (real n)" by simp
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   154
  next
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   155
    have "real n < real (2 * 2 ^ floor_log n)"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   156
      by (subst of_nat_less_iff) (rule floor_log_exp2_gt)
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   157
    also have "\<dots> = 2 powr (real (floor_log n) + 1)"
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   158
      by (simp add: powr_add powr_realpow)
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   159
    finally have "log 2 (real n) < log 2 \<dots>"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   160
      using False by (subst log_less_cancel_iff) simp_all
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   161
    also have "\<dots> = real (floor_log n) + 1" by simp
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   162
    finally show "log 2 (real n) < real_of_int (int (floor_log n)) + 1" by simp
64449
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   163
  qed
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   164
  thus ?thesis by simp
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   165
qed simp_all
8c44dfb4ca8a Merged natlog2 into Discrete.log
eberlm <eberlm@in.tum.de>
parents: 63766
diff changeset
   166
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   167
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 59349
diff changeset
   168
subsection \<open>Discrete square root\<close>
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   169
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   170
definition floor_sqrt :: "nat \<Rightarrow> nat"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   171
  where "floor_sqrt n = Max {m. m\<^sup>2 \<le> n}"
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   172
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   173
lemma floor_sqrt_aux:
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   174
  fixes n :: nat
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   175
  shows "finite {m. m\<^sup>2 \<le> n}" and "{m. m\<^sup>2 \<le> n} \<noteq> {}"
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   176
proof -
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 80621
diff changeset
   177
  have **: "m \<le> n" if "m\<^sup>2 \<le> n" for m
f94b30fa2b6c tuned proofs;
wenzelm
parents: 80621
diff changeset
   178
    using that by (cases m) (simp_all add: power2_eq_square)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   179
  then have "{m. m\<^sup>2 \<le> n} \<subseteq> {m. m \<le> n}" by auto
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   180
  then show "finite {m. m\<^sup>2 \<le> n}" by (rule finite_subset) rule
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   181
  have "0\<^sup>2 \<le> n" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   182
  then show *: "{m. m\<^sup>2 \<le> n} \<noteq> {}" by blast
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   183
qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   184
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   185
lemma floor_sqrt_unique:
63766
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   186
  assumes "m^2 \<le> n" "n < (Suc m)^2"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   187
  shows   "floor_sqrt n = m"
63766
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   188
proof -
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   189
  have "m' \<le> m" if "m'^2 \<le> n" for m'
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   190
  proof -
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   191
    note that
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   192
    also note assms(2)
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   193
    finally have "m' < Suc m" by (rule power_less_imp_less_base) simp_all
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   194
    thus "m' \<le> m" by simp
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   195
  qed
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   196
  with \<open>m^2 \<le> n\<close> floor_sqrt_aux[of n] show ?thesis unfolding floor_sqrt_def
63766
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   197
    by (intro antisym Max.boundedI Max.coboundedI) simp_all
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   198
qed
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   199
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   200
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   201
lemma floor_sqrt_code[code]: "floor_sqrt n = Max (Set.filter (\<lambda>m. m\<^sup>2 \<le> n) {0..n})"
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   202
proof -
82664
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 81467
diff changeset
   203
  from power2_nat_le_imp_le [of _ n]
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 81467
diff changeset
   204
    have "{m. m \<le> n \<and> m\<^sup>2 \<le> n} = {m. m\<^sup>2 \<le> n}" by auto
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 81467
diff changeset
   205
    then show ?thesis
e9f3b94eb6a0 more modern qualification of auxiliary operations
haftmann
parents: 81467
diff changeset
   206
    by (simp add: floor_sqrt_def)
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   207
qed
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   208
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   209
lemma floor_sqrt_inverse_power2 [simp]: "floor_sqrt (n\<^sup>2) = n"
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   210
proof -
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   211
  have "{m. m \<le> n} \<noteq> {}" by auto
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   212
  then have "Max {m. m \<le> n} \<le> n" by auto
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   213
  then show ?thesis
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   214
    by (auto simp add: floor_sqrt_def power2_nat_le_eq_le intro: antisym)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   215
qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   216
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   217
lemma floor_sqrt_zero [simp]: "floor_sqrt 0 = 0"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   218
  using floor_sqrt_inverse_power2 [of 0] by simp
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 57514
diff changeset
   219
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   220
lemma floor_sqrt_one [simp]: "floor_sqrt 1 = 1"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   221
  using floor_sqrt_inverse_power2 [of 1] by simp
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 57514
diff changeset
   222
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   223
lemma mono_floor_sqrt: "mono floor_sqrt"
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   224
proof
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   225
  fix m n :: nat
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   226
  have *: "0 * 0 \<le> m" by simp
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   227
  assume "m \<le> n"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   228
  then show "floor_sqrt m \<le> floor_sqrt n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   229
    by (auto intro!: Max_mono \<open>0 * 0 \<le> m\<close> finite_less_ub simp add: power2_eq_square floor_sqrt_def)
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   230
qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   231
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   232
lemma mono_floor_sqrt': "m \<le> n \<Longrightarrow> floor_sqrt m \<le> floor_sqrt n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   233
  using mono_floor_sqrt unfolding mono_def by auto
63766
695d60817cb1 Some facts about factorial and binomial coefficients
Manuel Eberl <eberlm@in.tum.de>
parents: 63540
diff changeset
   234
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   235
lemma floor_sqrt_greater_zero_iff [simp]: "floor_sqrt n > 0 \<longleftrightarrow> n > 0"
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   236
proof -
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   237
  have *: "0 < Max {m. m\<^sup>2 \<le> n} \<longleftrightarrow> (\<exists>a\<in>{m. m\<^sup>2 \<le> n}. 0 < a)"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   238
    by (rule Max_gr_iff) (fact floor_sqrt_aux)+
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   239
  show ?thesis
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   240
  proof
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   241
    assume "0 < floor_sqrt n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   242
    then have "0 < Max {m. m\<^sup>2 \<le> n}" by (simp add: floor_sqrt_def)
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   243
    with * show "0 < n" by (auto dest: power2_nat_le_imp_le)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   244
  next
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   245
    assume "0 < n"
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   246
    then have "1\<^sup>2 \<le> n \<and> 0 < (1::nat)" by simp
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   247
    then have "\<exists>q. q\<^sup>2 \<le> n \<and> 0 < q" ..
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   248
    with * have "0 < Max {m. m\<^sup>2 \<le> n}" by blast
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   249
    then show "0 < floor_sqrt n" by  (simp add: floor_sqrt_def)
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   250
  qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   251
qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   252
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   253
lemma floor_sqrt_power2_le [simp]: "(floor_sqrt n)\<^sup>2 \<le> n" (* FIXME tune proof *)
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   254
proof (cases "n > 0")
58787
af9eb5e566dd eliminated redundancies;
haftmann
parents: 57514
diff changeset
   255
  case False then show ?thesis by simp
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   256
next
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   257
  case True then have "floor_sqrt n > 0" by simp
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   258
  then have "mono (times (Max {m. m\<^sup>2 \<le> n}))" by (auto intro: mono_times_nat simp add: floor_sqrt_def)
53015
a1119cf551e8 standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents: 51263
diff changeset
   259
  then have *: "Max {m. m\<^sup>2 \<le> n} * Max {m. m\<^sup>2 \<le> n} = Max (times (Max {m. m\<^sup>2 \<le> n}) ` {m. m\<^sup>2 \<le> n})"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   260
    using floor_sqrt_aux [of n] by (rule mono_Max_commute)
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   261
  have "\<And>a. a * a \<le> n \<Longrightarrow> Max {m. m * m \<le> n} * a \<le> n"
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   262
  proof -
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   263
    fix q
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   264
    assume "q * q \<le> n"
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   265
    show "Max {m. m * m \<le> n} * q \<le> n"
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   266
    proof (cases "q > 0")
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   267
      case False then show ?thesis by simp
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   268
    next
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   269
      case True then have "mono (times q)" by (rule mono_times_nat)
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   270
      then have "q * Max {m. m * m \<le> n} = Max (times q ` {m. m * m \<le> n})"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   271
        using floor_sqrt_aux [of n] by (auto simp add: power2_eq_square intro: mono_Max_commute)
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   272
      then have "Max {m. m * m \<le> n} * q = Max (times q ` {m. m * m \<le> n})" by (simp add: ac_simps)
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67399
diff changeset
   273
      moreover have "finite ((*) q ` {m. m * m \<le> n})"
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   274
        by (metis (mono_tags) finite_imageI finite_less_ub le_square)
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   275
      moreover have "\<exists>x. x * x \<le> n"
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   276
        by (metis \<open>q * q \<le> n\<close>)
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   277
      ultimately show ?thesis
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   278
        by simp (metis \<open>q * q \<le> n\<close> le_cases mult_le_mono1 mult_le_mono2 order_trans)
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   279
    qed
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   280
  qed
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 67399
diff changeset
   281
  then have "Max ((*) (Max {m. m * m \<le> n}) ` {m. m * m \<le> n}) \<le> n"
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   282
    apply (subst Max_le_iff)
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   283
      apply (metis (mono_tags) finite_imageI finite_less_ub le_square)
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   284
     apply auto
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   285
    apply (metis le0 mult_0_right)
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   286
    done
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   287
  with * show ?thesis by (simp add: floor_sqrt_def power2_eq_square)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   288
qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   289
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   290
lemma floor_sqrt_le: "floor_sqrt n \<le> n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   291
  using floor_sqrt_aux [of n] by (auto simp add: floor_sqrt_def intro: power2_nat_le_imp_le)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   292
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   293
text \<open>Additional facts about the discrete square root, thanks to Julian Biendarra, Manuel Eberl\<close>
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 80621
diff changeset
   294
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   295
lemma Suc_floor_sqrt_power2_gt: "n < (Suc (floor_sqrt n))^2"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   296
  using Max_ge[OF floor_sqrt_aux(1), of "floor_sqrt n + 1" n]
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   297
  by (cases "n < (Suc (floor_sqrt n))^2") (simp_all add: floor_sqrt_def)
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   298
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   299
lemma le_floor_sqrt_iff: "x \<le> floor_sqrt y \<longleftrightarrow> x^2 \<le> y"
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   300
proof -
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   301
  have "x \<le> floor_sqrt y \<longleftrightarrow> (\<exists>z. z\<^sup>2 \<le> y \<and> x \<le> z)"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   302
    using Max_ge_iff[OF floor_sqrt_aux, of x y] by (simp add: floor_sqrt_def)
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   303
  also have "\<dots> \<longleftrightarrow> x^2 \<le> y"
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   304
  proof safe
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   305
    fix z assume "x \<le> z" "z ^ 2 \<le> y"
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   306
    thus "x^2 \<le> y" by (intro le_trans[of "x^2" "z^2" y]) (simp_all add: power2_nat_le_eq_le)
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   307
  qed auto
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   308
  finally show ?thesis .
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   309
qed
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 80621
diff changeset
   310
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   311
lemma le_floor_sqrtI: "x^2 \<le> y \<Longrightarrow> x \<le> floor_sqrt y"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   312
  by (simp add: le_floor_sqrt_iff)
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   313
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   314
lemma floor_sqrt_le_iff: "floor_sqrt y \<le> x \<longleftrightarrow> (\<forall>z. z^2 \<le> y \<longrightarrow> z \<le> x)"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   315
  using Max.bounded_iff[OF floor_sqrt_aux] by (simp add: floor_sqrt_def)
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   316
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   317
lemma floor_sqrt_leI:
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   318
  "(\<And>z. z^2 \<le> y \<Longrightarrow> z \<le> x) \<Longrightarrow> floor_sqrt y \<le> x"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   319
  by (simp add: floor_sqrt_le_iff)
81332
f94b30fa2b6c tuned proofs;
wenzelm
parents: 80621
diff changeset
   320
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   321
lemma floor_sqrt_Suc:
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   322
  "floor_sqrt (Suc n) = (if \<exists>m. Suc n = m^2 then Suc (floor_sqrt n) else floor_sqrt n)"
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   323
proof cases
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   324
  assume "\<exists> m. Suc n = m^2"
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   325
  then obtain m where m_def: "Suc n = m^2" by blast
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   326
  then have lhs: "floor_sqrt (Suc n) = m" by simp
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   327
  from m_def floor_sqrt_power2_le[of n]
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   328
    have "(floor_sqrt n)^2 < m^2" by linarith
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   329
  with power2_less_imp_less have lt_m: "floor_sqrt n < m" by blast
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   330
  from m_def Suc_floor_sqrt_power2_gt[of "n"]
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   331
    have "m^2 \<le> (Suc(floor_sqrt n))^2"
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 69064
diff changeset
   332
      by linarith
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   333
  with power2_nat_le_eq_le have "m \<le> Suc (floor_sqrt n)" by blast
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   334
  with lt_m have "m = Suc (floor_sqrt n)" by simp
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   335
  with lhs m_def show ?thesis by fastforce
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   336
next
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   337
  assume asm: "\<not> (\<exists> m. Suc n = m^2)"
81467
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   338
  hence "Suc n \<noteq> (floor_sqrt (Suc n))^2" by simp
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   339
  with floor_sqrt_power2_le[of "Suc n"]
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   340
    have "floor_sqrt (Suc n) \<le> floor_sqrt n" by (intro le_floor_sqrtI) linarith
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   341
  moreover have "floor_sqrt (Suc n) \<ge> floor_sqrt n"
3fab5b28027d renamed Discrete -> Discrete_Functions to avoid name clashes;
nipkow
parents: 81332
diff changeset
   342
    by (intro monoD[OF mono_floor_sqrt]) simp_all
64919
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   343
  ultimately show ?thesis using asm by simp
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   344
qed
7e0c8924dfda New material from Bertrand's postulate (file Bertrand_Discrete_Sqrt), and tidied an old messy proof
paulson <lp15@cam.ac.uk>
parents: 64449
diff changeset
   345
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   346
end