src/HOL/Library/Discrete.thy
author wenzelm
Wed, 07 Aug 2013 14:47:50 +0200
changeset 52891 b8dede3a4f1d
parent 51263 31e786e0e6a7
child 53015 a1119cf551e8
permissions -rw-r--r--
tuned proofs;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)  
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     2
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     3
header {* Common discrete functions *}
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     4
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     5
theory Discrete
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     6
imports Main
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
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
     9
subsection {* Discrete logarithm *}
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    10
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    11
fun log :: "nat \<Rightarrow> nat" where
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    12
  [simp del]: "log n = (if n < 2 then 0 else Suc (log (n div 2)))"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    13
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    14
lemma log_zero [simp]:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    15
  "log 0 = 0"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    16
  by (simp add: log.simps)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    17
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    18
lemma log_one [simp]:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    19
  "log 1 = 0"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    20
  by (simp add: log.simps)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    21
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    22
lemma log_Suc_zero [simp]:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    23
  "log (Suc 0) = 0"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    24
  using log_one by simp
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    25
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    26
lemma log_rec:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    27
  "n \<ge> 2 \<Longrightarrow> log n = Suc (log (n div 2))"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    28
  by (simp add: log.simps)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    29
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    30
lemma log_twice [simp]:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    31
  "n \<noteq> 0 \<Longrightarrow> log (2 * n) = Suc (log n)"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    32
  by (simp add: log_rec)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    33
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    34
lemma log_half [simp]:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    35
  "log (n div 2) = log n - 1"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    36
proof (cases "n < 2")
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    37
  case True
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    38
  then have "n = 0 \<or> n = 1" by arith
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    39
  then show ?thesis by (auto simp del: One_nat_def)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    40
next
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    41
  case False then show ?thesis by (simp add: log_rec)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    42
qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    43
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    44
lemma log_exp [simp]:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    45
  "log (2 ^ n) = n"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    46
  by (induct n) simp_all
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    47
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    48
lemma log_mono:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    49
  "mono log"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    50
proof
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    51
  fix m n :: nat
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    52
  assume "m \<le> n"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    53
  then show "log m \<le> log n"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    54
  proof (induct m arbitrary: n rule: log.induct)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    55
    case (1 m)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    56
    then have mn2: "m div 2 \<le> n div 2" by arith
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    57
    show "log m \<le> log n"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    58
    proof (cases "m < 2")
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    59
      case True
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    60
      then have "m = 0 \<or> m = 1" by arith
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    61
      then show ?thesis by (auto simp del: One_nat_def)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    62
    next
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    63
      case False
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    64
      with mn2 have "m \<ge> 2" and "n \<ge> 2" by auto arith
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    65
      from False have m2_0: "m div 2 \<noteq> 0" by arith
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    66
      with mn2 have n2_0: "n div 2 \<noteq> 0" by arith
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    67
      from False "1.hyps" mn2 have "log (m div 2) \<le> log (n div 2)" by blast
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    68
      with m2_0 n2_0 have "log (2 * (m div 2)) \<le> log (2 * (n div 2))" by simp
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    69
      with m2_0 n2_0 `m \<ge> 2` `n \<ge> 2` show ?thesis by (simp only: log_rec [of m] log_rec [of n]) simp
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    70
    qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    71
  qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    72
qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    73
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    74
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    75
subsection {* Discrete square root *}
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    76
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    77
definition sqrt :: "nat \<Rightarrow> nat"
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
    78
where
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    79
  "sqrt n = Max {m. m\<twosuperior> \<le> n}"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    80
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    81
lemma sqrt_aux:
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    82
  fixes n :: nat
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    83
  shows "finite {m. m\<twosuperior> \<le> n}" and "{m. m\<twosuperior> \<le> n} \<noteq> {}"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    84
proof -
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    85
  { fix m
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    86
    assume "m\<twosuperior> \<le> n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    87
    then have "m \<le> n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    88
      by (cases m) (simp_all add: power2_eq_square)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    89
  } note ** = this
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    90
  then have "{m. m\<twosuperior> \<le> n} \<subseteq> {m. m \<le> n}" by auto
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    91
  then show "finite {m. m\<twosuperior> \<le> n}" by (rule finite_subset) rule
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    92
  have "0\<twosuperior> \<le> n" by simp
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    93
  then show *: "{m. m\<twosuperior> \<le> n} \<noteq> {}" by blast
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    94
qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    95
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    96
lemma [code]:
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    97
  "sqrt n = Max (Set.filter (\<lambda>m. m\<twosuperior> \<le> n) {0..n})"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    98
proof -
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
    99
  from power2_nat_le_imp_le [of _ n] have "{m. m \<le> n \<and> m\<twosuperior> \<le> n} = {m. m\<twosuperior> \<le> n}" by auto
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   100
  then show ?thesis by (simp add: sqrt_def Set.filter_def)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   101
qed
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   102
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   103
lemma sqrt_inverse_power2 [simp]:
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   104
  "sqrt (n\<twosuperior>) = n"
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   105
proof -
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   106
  have "{m. m \<le> n} \<noteq> {}" by auto
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   107
  then have "Max {m. m \<le> n} \<le> n" by auto
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   108
  then show ?thesis
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   109
    by (auto simp add: sqrt_def power2_nat_le_eq_le intro: antisym)
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   110
qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   111
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   112
lemma mono_sqrt: "mono sqrt"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   113
proof
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   114
  fix m n :: nat
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   115
  have *: "0 * 0 \<le> m" by simp
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   116
  assume "m \<le> n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   117
  then show "sqrt m \<le> sqrt n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   118
    by (auto intro!: Max_mono `0 * 0 \<le> m` finite_less_ub simp add: power2_eq_square sqrt_def)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   119
qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   120
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   121
lemma sqrt_greater_zero_iff [simp]:
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   122
  "sqrt n > 0 \<longleftrightarrow> n > 0"
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   123
proof -
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   124
  have *: "0 < Max {m. m\<twosuperior> \<le> n} \<longleftrightarrow> (\<exists>a\<in>{m. m\<twosuperior> \<le> n}. 0 < a)"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   125
    by (rule Max_gr_iff) (fact sqrt_aux)+
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   126
  show ?thesis
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   127
  proof
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   128
    assume "0 < sqrt n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   129
    then have "0 < Max {m. m\<twosuperior> \<le> n}" by (simp add: sqrt_def)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   130
    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
   131
  next
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   132
    assume "0 < n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   133
    then have "1\<twosuperior> \<le> n \<and> 0 < (1::nat)" by simp
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   134
    then have "\<exists>q. q\<twosuperior> \<le> n \<and> 0 < q" ..
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   135
    with * have "0 < Max {m. m\<twosuperior> \<le> n}" by blast
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   136
    then show "0 < sqrt n" by  (simp add: sqrt_def)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   137
  qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   138
qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   139
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   140
lemma sqrt_power2_le [simp]: (* FIXME tune proof *)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   141
  "(sqrt n)\<twosuperior> \<le> n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   142
proof (cases "n > 0")
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   143
  case False then show ?thesis by (simp add: sqrt_def)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   144
next
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   145
  case True then have "sqrt n > 0" by simp
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   146
  then have "mono (times (Max {m. m\<twosuperior> \<le> n}))" by (auto intro: mono_times_nat simp add: sqrt_def)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   147
  then have *: "Max {m. m\<twosuperior> \<le> n} * Max {m. m\<twosuperior> \<le> n} = Max (times (Max {m. m\<twosuperior> \<le> n}) ` {m. m\<twosuperior> \<le> n})"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   148
    using sqrt_aux [of n] by (rule mono_Max_commute)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   149
  have "Max (op * (Max {m. m * m \<le> n}) ` {m. m * m \<le> n}) \<le> n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   150
    apply (subst Max_le_iff)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   151
    apply (metis (mono_tags) finite_imageI finite_less_ub le_square)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   152
    apply simp
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   153
    apply (metis le0 mult_0_right)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   154
    apply auto
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   155
    proof -
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   156
      fix q
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   157
      assume "q * q \<le> n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   158
      show "Max {m. m * m \<le> n} * q \<le> n"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   159
      proof (cases "q > 0")
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   160
        case False then show ?thesis by simp
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   161
      next
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   162
        case True then have "mono (times q)" by (rule mono_times_nat)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   163
        then have "q * Max {m. m * m \<le> n} = Max (times q ` {m. m * m \<le> n})"
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   164
          using sqrt_aux [of n] by (auto simp add: power2_eq_square intro: mono_Max_commute)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   165
        then have "Max {m. m * m \<le> n} * q = Max (times q ` {m. m * m \<le> n})" by (simp add: mult_ac)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   166
        then show ?thesis apply simp
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   167
          apply (subst Max_le_iff)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   168
          apply auto
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   169
          apply (metis (mono_tags) finite_imageI finite_less_ub le_square)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   170
          apply (metis `q * q \<le> n`)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   171
          using `q * q \<le> n` by (metis le_cases mult_le_mono1 mult_le_mono2 order_trans)
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   172
      qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   173
    qed
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   174
  with * show ?thesis by (simp add: sqrt_def power2_eq_square)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   175
qed
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   176
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   177
lemma sqrt_le:
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   178
  "sqrt n \<le> n"
51263
31e786e0e6a7 turned example into library for comparing growth of functions
haftmann
parents: 51174
diff changeset
   179
  using sqrt_aux [of n] by (auto simp add: sqrt_def intro: power2_nat_le_imp_le)
51174
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   180
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   181
hide_const (open) log sqrt
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   182
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   183
end
071674018df9 fundamentals about discrete logarithm and square root
haftmann
parents:
diff changeset
   184