src/HOL/NthRoot.thy
author blanchet
Thu, 14 Apr 2011 11:24:05 +0200
changeset 42351 ad89f5462cdc
parent 35216 7641e8d831d2
child 44289 d81d09cdab9c
permissions -rw-r--r--
remove needless export
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     1
(*  Title       : NthRoot.thy
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     2
    Author      : Jacques D. Fleuriot
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     3
    Copyright   : 1998  University of Cambridge
14477
cc61fd03e589 conversion of Hyperreal/Lim to new-style
paulson
parents: 14365
diff changeset
     4
    Conversion to Isar and new proofs by Lawrence C Paulson, 2004
12196
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     5
*)
a3be6b3a9c0b new theories from Jacques Fleuriot
paulson
parents:
diff changeset
     6
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
     7
header {* Nth Roots of Real Numbers *}
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
     8
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
     9
theory NthRoot
28952
15a4b2cf8c34 made repository layout more coherent with logical distribution structure; stripped some $Id$s
haftmann
parents: 25875
diff changeset
    10
imports Parity Deriv
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15085
diff changeset
    11
begin
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    12
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    13
subsection {* Existence of Nth Root *}
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    14
23009
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    15
text {* Existence follows from the Intermediate Value Theorem *}
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    16
23009
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    17
lemma realpow_pos_nth:
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    18
  assumes n: "0 < n"
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    19
  assumes a: "0 < a"
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    20
  shows "\<exists>r>0. r ^ n = (a::real)"
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    21
proof -
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    22
  have "\<exists>r\<ge>0. r \<le> (max 1 a) \<and> r ^ n = a"
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    23
  proof (rule IVT)
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    24
    show "0 ^ n \<le> a" using n a by (simp add: power_0_left)
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    25
    show "0 \<le> max 1 a" by simp
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    26
    from n have n1: "1 \<le> n" by simp
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    27
    have "a \<le> max 1 a ^ 1" by simp
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    28
    also have "max 1 a ^ 1 \<le> max 1 a ^ n"
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    29
      using n1 by (rule power_increasing, simp)
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    30
    finally show "a \<le> max 1 a ^ n" .
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    31
    show "\<forall>r. 0 \<le> r \<and> r \<le> max 1 a \<longrightarrow> isCont (\<lambda>x. x ^ n) r"
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 23049
diff changeset
    32
      by (simp add: isCont_power)
23009
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    33
  qed
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    34
  then obtain r where r: "0 \<le> r \<and> r ^ n = a" by fast
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    35
  with n a have "r \<noteq> 0" by (auto simp add: power_0_left)
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    36
  with r have "0 < r \<and> r ^ n = a" by simp
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    37
  thus ?thesis ..
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    38
qed
14325
94ac3895822f removing real_of_posnat
paulson
parents: 14324
diff changeset
    39
23047
17f7d831efe2 add realpow_pos_nth2 back in
huffman
parents: 23046
diff changeset
    40
(* Used by Integration/RealRandVar.thy in AFP *)
17f7d831efe2 add realpow_pos_nth2 back in
huffman
parents: 23046
diff changeset
    41
lemma realpow_pos_nth2: "(0::real) < a \<Longrightarrow> \<exists>r>0. r ^ Suc n = a"
17f7d831efe2 add realpow_pos_nth2 back in
huffman
parents: 23046
diff changeset
    42
by (blast intro: realpow_pos_nth)
17f7d831efe2 add realpow_pos_nth2 back in
huffman
parents: 23046
diff changeset
    43
23009
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    44
text {* Uniqueness of nth positive root *}
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    45
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    46
lemma realpow_pos_nth_unique:
23009
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    47
  "\<lbrakk>0 < n; 0 < a\<rbrakk> \<Longrightarrow> \<exists>!r. 0 < r \<and> r ^ n = (a::real)"
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    48
apply (auto intro!: realpow_pos_nth)
23009
01c295dd4a36 Prove existence of nth roots using Intermediate Value Theorem
huffman
parents: 22968
diff changeset
    49
apply (rule_tac n=n in power_eq_imp_eq_base, simp_all)
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    50
done
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
    51
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    52
subsection {* Nth Root *}
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    53
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    54
text {* We define roots of negative reals such that
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    55
  @{term "root n (- x) = - root n x"}. This allows
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    56
  us to omit side conditions from many theorems. *}
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    57
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    58
definition
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    59
  root :: "[nat, real] \<Rightarrow> real" where
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    60
  "root n x = (if 0 < x then (THE u. 0 < u \<and> u ^ n = x) else
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    61
               if x < 0 then - (THE u. 0 < u \<and> u ^ n = - x) else 0)"
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    62
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    63
lemma real_root_zero [simp]: "root n 0 = 0"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    64
unfolding root_def by simp
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    65
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    66
lemma real_root_minus: "0 < n \<Longrightarrow> root n (- x) = - root n x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    67
unfolding root_def by simp
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    68
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    69
lemma real_root_gt_zero: "\<lbrakk>0 < n; 0 < x\<rbrakk> \<Longrightarrow> 0 < root n x"
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    70
apply (simp add: root_def)
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    71
apply (drule (1) realpow_pos_nth_unique)
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    72
apply (erule theI' [THEN conjunct1])
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    73
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    74
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    75
lemma real_root_pow_pos: (* TODO: rename *)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    76
  "\<lbrakk>0 < n; 0 < x\<rbrakk> \<Longrightarrow> root n x ^ n = x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    77
apply (simp add: root_def)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    78
apply (drule (1) realpow_pos_nth_unique)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    79
apply (erule theI' [THEN conjunct2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    80
done
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    81
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    82
lemma real_root_pow_pos2 [simp]: (* TODO: rename *)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    83
  "\<lbrakk>0 < n; 0 \<le> x\<rbrakk> \<Longrightarrow> root n x ^ n = x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    84
by (auto simp add: order_le_less real_root_pow_pos)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    85
23046
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
    86
lemma odd_real_root_pow: "odd n \<Longrightarrow> root n x ^ n = x"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
    87
apply (rule_tac x=0 and y=x in linorder_le_cases)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
    88
apply (erule (1) real_root_pow_pos2 [OF odd_pos])
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
    89
apply (subgoal_tac "root n (- x) ^ n = - x")
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
    90
apply (simp add: real_root_minus odd_pos)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
    91
apply (simp add: odd_pos)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
    92
done
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
    93
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    94
lemma real_root_ge_zero: "\<lbrakk>0 < n; 0 \<le> x\<rbrakk> \<Longrightarrow> 0 \<le> root n x"
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    95
by (auto simp add: order_le_less real_root_gt_zero)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
    96
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    97
lemma real_root_power_cancel: "\<lbrakk>0 < n; 0 \<le> x\<rbrakk> \<Longrightarrow> root n (x ^ n) = x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    98
apply (subgoal_tac "0 \<le> x ^ n")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
    99
apply (subgoal_tac "0 \<le> root n (x ^ n)")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   100
apply (subgoal_tac "root n (x ^ n) ^ n = x ^ n")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   101
apply (erule (3) power_eq_imp_eq_base)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   102
apply (erule (1) real_root_pow_pos2)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   103
apply (erule (1) real_root_ge_zero)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   104
apply (erule zero_le_power)
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   105
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   106
23046
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   107
lemma odd_real_root_power_cancel: "odd n \<Longrightarrow> root n (x ^ n) = x"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   108
apply (rule_tac x=0 and y=x in linorder_le_cases)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   109
apply (erule (1) real_root_power_cancel [OF odd_pos])
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   110
apply (subgoal_tac "root n ((- x) ^ n) = - x")
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   111
apply (simp add: real_root_minus odd_pos)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   112
apply (erule real_root_power_cancel [OF odd_pos], simp)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   113
done
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   114
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   115
lemma real_root_pos_unique:
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   116
  "\<lbrakk>0 < n; 0 \<le> y; y ^ n = x\<rbrakk> \<Longrightarrow> root n x = y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   117
by (erule subst, rule real_root_power_cancel)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   118
23046
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   119
lemma odd_real_root_unique:
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   120
  "\<lbrakk>odd n; y ^ n = x\<rbrakk> \<Longrightarrow> root n x = y"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   121
by (erule subst, rule odd_real_root_power_cancel)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   122
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   123
lemma real_root_one [simp]: "0 < n \<Longrightarrow> root n 1 = 1"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   124
by (simp add: real_root_pos_unique)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   125
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   126
text {* Root function is strictly monotonic, hence injective *}
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   127
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   128
lemma real_root_less_mono_lemma:
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   129
  "\<lbrakk>0 < n; 0 \<le> x; x < y\<rbrakk> \<Longrightarrow> root n x < root n y"
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   130
apply (subgoal_tac "0 \<le> y")
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   131
apply (subgoal_tac "root n x ^ n < root n y ^ n")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   132
apply (erule power_less_imp_less_base)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   133
apply (erule (1) real_root_ge_zero)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   134
apply simp
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   135
apply simp
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   136
done
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   137
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   138
lemma real_root_less_mono: "\<lbrakk>0 < n; x < y\<rbrakk> \<Longrightarrow> root n x < root n y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   139
apply (cases "0 \<le> x")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   140
apply (erule (2) real_root_less_mono_lemma)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   141
apply (cases "0 \<le> y")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   142
apply (rule_tac y=0 in order_less_le_trans)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   143
apply (subgoal_tac "0 < root n (- x)")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   144
apply (simp add: real_root_minus)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   145
apply (simp add: real_root_gt_zero)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   146
apply (simp add: real_root_ge_zero)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   147
apply (subgoal_tac "root n (- y) < root n (- x)")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   148
apply (simp add: real_root_minus)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   149
apply (simp add: real_root_less_mono_lemma)
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   150
done
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   151
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   152
lemma real_root_le_mono: "\<lbrakk>0 < n; x \<le> y\<rbrakk> \<Longrightarrow> root n x \<le> root n y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   153
by (auto simp add: order_le_less real_root_less_mono)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   154
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   155
lemma real_root_less_iff [simp]:
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   156
  "0 < n \<Longrightarrow> (root n x < root n y) = (x < y)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   157
apply (cases "x < y")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   158
apply (simp add: real_root_less_mono)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   159
apply (simp add: linorder_not_less real_root_le_mono)
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   160
done
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   161
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   162
lemma real_root_le_iff [simp]:
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   163
  "0 < n \<Longrightarrow> (root n x \<le> root n y) = (x \<le> y)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   164
apply (cases "x \<le> y")
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   165
apply (simp add: real_root_le_mono)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   166
apply (simp add: linorder_not_le real_root_less_mono)
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   167
done
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   168
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   169
lemma real_root_eq_iff [simp]:
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   170
  "0 < n \<Longrightarrow> (root n x = root n y) = (x = y)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   171
by (simp add: order_eq_iff)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   172
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   173
lemmas real_root_gt_0_iff [simp] = real_root_less_iff [where x=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   174
lemmas real_root_lt_0_iff [simp] = real_root_less_iff [where y=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   175
lemmas real_root_ge_0_iff [simp] = real_root_le_iff [where x=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   176
lemmas real_root_le_0_iff [simp] = real_root_le_iff [where y=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   177
lemmas real_root_eq_0_iff [simp] = real_root_eq_iff [where y=0, simplified]
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   178
23257
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   179
lemma real_root_gt_1_iff [simp]: "0 < n \<Longrightarrow> (1 < root n y) = (1 < y)"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   180
by (insert real_root_less_iff [where x=1], simp)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   181
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   182
lemma real_root_lt_1_iff [simp]: "0 < n \<Longrightarrow> (root n x < 1) = (x < 1)"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   183
by (insert real_root_less_iff [where y=1], simp)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   184
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   185
lemma real_root_ge_1_iff [simp]: "0 < n \<Longrightarrow> (1 \<le> root n y) = (1 \<le> y)"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   186
by (insert real_root_le_iff [where x=1], simp)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   187
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   188
lemma real_root_le_1_iff [simp]: "0 < n \<Longrightarrow> (root n x \<le> 1) = (x \<le> 1)"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   189
by (insert real_root_le_iff [where y=1], simp)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   190
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   191
lemma real_root_eq_1_iff [simp]: "0 < n \<Longrightarrow> (root n x = 1) = (x = 1)"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   192
by (insert real_root_eq_iff [where y=1], simp)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   193
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   194
text {* Roots of roots *}
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   195
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   196
lemma real_root_Suc_0 [simp]: "root (Suc 0) x = x"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   197
by (simp add: odd_real_root_unique)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   198
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   199
lemma real_root_pos_mult_exp:
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   200
  "\<lbrakk>0 < m; 0 < n; 0 < x\<rbrakk> \<Longrightarrow> root (m * n) x = root m (root n x)"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   201
by (rule real_root_pos_unique, simp_all add: power_mult)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   202
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   203
lemma real_root_mult_exp:
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   204
  "\<lbrakk>0 < m; 0 < n\<rbrakk> \<Longrightarrow> root (m * n) x = root m (root n x)"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   205
apply (rule linorder_cases [where x=x and y=0])
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   206
apply (subgoal_tac "root (m * n) (- x) = root m (root n (- x))")
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   207
apply (simp add: real_root_minus)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   208
apply (simp_all add: real_root_pos_mult_exp)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   209
done
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   210
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   211
lemma real_root_commute:
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   212
  "\<lbrakk>0 < m; 0 < n\<rbrakk> \<Longrightarrow> root m (root n x) = root n (root m x)"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   213
by (simp add: real_root_mult_exp [symmetric] mult_commute)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   214
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   215
text {* Monotonicity in first argument *}
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   216
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   217
lemma real_root_strict_decreasing:
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   218
  "\<lbrakk>0 < n; n < N; 1 < x\<rbrakk> \<Longrightarrow> root N x < root n x"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   219
apply (subgoal_tac "root n (root N x) ^ n < root N (root n x) ^ N", simp)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   220
apply (simp add: real_root_commute power_strict_increasing
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   221
            del: real_root_pow_pos2)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   222
done
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   223
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   224
lemma real_root_strict_increasing:
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   225
  "\<lbrakk>0 < n; n < N; 0 < x; x < 1\<rbrakk> \<Longrightarrow> root n x < root N x"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   226
apply (subgoal_tac "root N (root n x) ^ N < root n (root N x) ^ n", simp)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   227
apply (simp add: real_root_commute power_strict_decreasing
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   228
            del: real_root_pow_pos2)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   229
done
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   230
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   231
lemma real_root_decreasing:
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   232
  "\<lbrakk>0 < n; n < N; 1 \<le> x\<rbrakk> \<Longrightarrow> root N x \<le> root n x"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   233
by (auto simp add: order_le_less real_root_strict_decreasing)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   234
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   235
lemma real_root_increasing:
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   236
  "\<lbrakk>0 < n; n < N; 0 \<le> x; x \<le> 1\<rbrakk> \<Longrightarrow> root n x \<le> root N x"
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   237
by (auto simp add: order_le_less real_root_strict_increasing)
9117e228a8e3 add new lemmas
huffman
parents: 23122
diff changeset
   238
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   239
text {* Roots of multiplication and division *}
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   240
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   241
lemma real_root_mult_lemma:
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   242
  "\<lbrakk>0 < n; 0 \<le> x; 0 \<le> y\<rbrakk> \<Longrightarrow> root n (x * y) = root n x * root n y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   243
by (simp add: real_root_pos_unique mult_nonneg_nonneg power_mult_distrib)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   244
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   245
lemma real_root_inverse_lemma:
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   246
  "\<lbrakk>0 < n; 0 \<le> x\<rbrakk> \<Longrightarrow> root n (inverse x) = inverse (root n x)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   247
by (simp add: real_root_pos_unique power_inverse [symmetric])
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   248
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   249
lemma real_root_mult:
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   250
  assumes n: "0 < n"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   251
  shows "root n (x * y) = root n x * root n y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   252
proof (rule linorder_le_cases, rule_tac [!] linorder_le_cases)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   253
  assume "0 \<le> x" and "0 \<le> y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   254
  thus ?thesis by (rule real_root_mult_lemma [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   255
next
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   256
  assume "0 \<le> x" and "y \<le> 0"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   257
  hence "0 \<le> x" and "0 \<le> - y" by simp_all
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   258
  hence "root n (x * - y) = root n x * root n (- y)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   259
    by (rule real_root_mult_lemma [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   260
  thus ?thesis by (simp add: real_root_minus [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   261
next
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   262
  assume "x \<le> 0" and "0 \<le> y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   263
  hence "0 \<le> - x" and "0 \<le> y" by simp_all
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   264
  hence "root n (- x * y) = root n (- x) * root n y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   265
    by (rule real_root_mult_lemma [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   266
  thus ?thesis by (simp add: real_root_minus [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   267
next
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   268
  assume "x \<le> 0" and "y \<le> 0"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   269
  hence "0 \<le> - x" and "0 \<le> - y" by simp_all
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   270
  hence "root n (- x * - y) = root n (- x) * root n (- y)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   271
    by (rule real_root_mult_lemma [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   272
  thus ?thesis by (simp add: real_root_minus [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   273
qed
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   274
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   275
lemma real_root_inverse:
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   276
  assumes n: "0 < n"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   277
  shows "root n (inverse x) = inverse (root n x)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   278
proof (rule linorder_le_cases)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   279
  assume "0 \<le> x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   280
  thus ?thesis by (rule real_root_inverse_lemma [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   281
next
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   282
  assume "x \<le> 0"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   283
  hence "0 \<le> - x" by simp
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   284
  hence "root n (inverse (- x)) = inverse (root n (- x))"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   285
    by (rule real_root_inverse_lemma [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   286
  thus ?thesis by (simp add: real_root_minus [OF n])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   287
qed
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   288
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   289
lemma real_root_divide:
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   290
  "0 < n \<Longrightarrow> root n (x / y) = root n x / root n y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   291
by (simp add: divide_inverse real_root_mult real_root_inverse)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   292
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   293
lemma real_root_power:
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   294
  "0 < n \<Longrightarrow> root n (x ^ k) = root n x ^ k"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   295
by (induct k, simp_all add: real_root_mult)
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   296
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   297
lemma real_root_abs: "0 < n \<Longrightarrow> root n \<bar>x\<bar> = \<bar>root n x\<bar>"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   298
by (simp add: abs_if real_root_minus)
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   299
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   300
text {* Continuity and derivatives *}
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   301
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   302
lemma isCont_root_pos:
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   303
  assumes n: "0 < n"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   304
  assumes x: "0 < x"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   305
  shows "isCont (root n) x"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   306
proof -
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   307
  have "isCont (root n) (root n x ^ n)"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   308
  proof (rule isCont_inverse_function [where f="\<lambda>a. a ^ n"])
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   309
    show "0 < root n x" using n x by simp
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   310
    show "\<forall>z. \<bar>z - root n x\<bar> \<le> root n x \<longrightarrow> root n (z ^ n) = z"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   311
      by (simp add: abs_le_iff real_root_power_cancel n)
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   312
    show "\<forall>z. \<bar>z - root n x\<bar> \<le> root n x \<longrightarrow> isCont (\<lambda>a. a ^ n) z"
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 23049
diff changeset
   313
      by (simp add: isCont_power)
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   314
  qed
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   315
  thus ?thesis using n x by simp
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   316
qed
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   317
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   318
lemma isCont_root_neg:
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   319
  "\<lbrakk>0 < n; x < 0\<rbrakk> \<Longrightarrow> isCont (root n) x"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   320
apply (subgoal_tac "isCont (\<lambda>x. - root n (- x)) x")
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   321
apply (simp add: real_root_minus)
23069
cdfff0241c12 rename lemmas LIM_ident, isCont_ident, DERIV_ident
huffman
parents: 23049
diff changeset
   322
apply (rule isCont_o2 [OF isCont_minus [OF isCont_ident]])
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   323
apply (simp add: isCont_minus isCont_root_pos)
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   324
done
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   325
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   326
lemma isCont_root_zero:
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   327
  "0 < n \<Longrightarrow> isCont (root n) 0"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   328
unfolding isCont_def
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   329
apply (rule LIM_I)
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   330
apply (rule_tac x="r ^ n" in exI, safe)
25875
536dfdc25e0a added simp attributes/ proofs fixed
nipkow
parents: 25766
diff changeset
   331
apply (simp)
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   332
apply (simp add: real_root_abs [symmetric])
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   333
apply (rule_tac n="n" in power_less_imp_less_base, simp_all)
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   334
done
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   335
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   336
lemma isCont_real_root: "0 < n \<Longrightarrow> isCont (root n) x"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   337
apply (rule_tac x=x and y=0 in linorder_cases)
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   338
apply (simp_all add: isCont_root_pos isCont_root_neg isCont_root_zero)
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   339
done
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   340
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   341
lemma DERIV_real_root:
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   342
  assumes n: "0 < n"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   343
  assumes x: "0 < x"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   344
  shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   345
proof (rule DERIV_inverse_function)
23044
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23042
diff changeset
   346
  show "0 < x" using x .
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23042
diff changeset
   347
  show "x < x + 1" by simp
2ad82c359175 change premises of DERIV_inverse_function lemma
huffman
parents: 23042
diff changeset
   348
  show "\<forall>y. 0 < y \<and> y < x + 1 \<longrightarrow> root n y ^ n = y"
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   349
    using n by simp
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   350
  show "DERIV (\<lambda>x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   351
    by (rule DERIV_pow)
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   352
  show "real n * root n x ^ (n - Suc 0) \<noteq> 0"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   353
    using n x by simp
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   354
  show "isCont (root n) x"
23441
ee218296d635 avoid using implicit prems in assumption
huffman
parents: 23257
diff changeset
   355
    using n by (rule isCont_real_root)
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   356
qed
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   357
23046
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   358
lemma DERIV_odd_real_root:
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   359
  assumes n: "odd n"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   360
  assumes x: "x \<noteq> 0"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   361
  shows "DERIV (root n) x :> inverse (real n * root n x ^ (n - Suc 0))"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   362
proof (rule DERIV_inverse_function)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   363
  show "x - 1 < x" by simp
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   364
  show "x < x + 1" by simp
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   365
  show "\<forall>y. x - 1 < y \<and> y < x + 1 \<longrightarrow> root n y ^ n = y"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   366
    using n by (simp add: odd_real_root_pow)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   367
  show "DERIV (\<lambda>x. x ^ n) (root n x) :> real n * root n x ^ (n - Suc 0)"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   368
    by (rule DERIV_pow)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   369
  show "real n * root n x ^ (n - Suc 0) \<noteq> 0"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   370
    using odd_pos [OF n] x by simp
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   371
  show "isCont (root n) x"
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   372
    using odd_pos [OF n] by (rule isCont_real_root)
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   373
qed
12f35ece221f add odd_real_root lemmas
huffman
parents: 23044
diff changeset
   374
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   375
lemma DERIV_even_real_root:
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   376
  assumes n: "0 < n" and "even n"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   377
  assumes x: "x < 0"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   378
  shows "DERIV (root n) x :> inverse (- real n * root n x ^ (n - Suc 0))"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   379
proof (rule DERIV_inverse_function)
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   380
  show "x - 1 < x" by simp
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   381
  show "x < 0" using x .
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   382
next
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   383
  show "\<forall>y. x - 1 < y \<and> y < 0 \<longrightarrow> - (root n y ^ n) = y"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   384
  proof (rule allI, rule impI, erule conjE)
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   385
    fix y assume "x - 1 < y" and "y < 0"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   386
    hence "root n (-y) ^ n = -y" using `0 < n` by simp
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   387
    with real_root_minus[OF `0 < n`] and `even n`
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   388
    show "- (root n y ^ n) = y" by simp
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   389
  qed
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   390
next
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   391
  show "DERIV (\<lambda>x. - (x ^ n)) (root n x) :> - real n * root n x ^ (n - Suc 0)"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   392
    by  (auto intro!: DERIV_intros)
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   393
  show "- real n * root n x ^ (n - Suc 0) \<noteq> 0"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   394
    using n x by simp
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   395
  show "isCont (root n) x"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   396
    using n by (rule isCont_real_root)
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   397
qed
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   398
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   399
lemma DERIV_real_root_generic:
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   400
  assumes "0 < n" and "x \<noteq> 0"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   401
  and even: "\<lbrakk> even n ; 0 < x \<rbrakk> \<Longrightarrow> D = inverse (real n * root n x ^ (n - Suc 0))"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   402
  and even: "\<lbrakk> even n ; x < 0 \<rbrakk> \<Longrightarrow> D = - inverse (real n * root n x ^ (n - Suc 0))"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   403
  and odd: "odd n \<Longrightarrow> D = inverse (real n * root n x ^ (n - Suc 0))"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   404
  shows "DERIV (root n) x :> D"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   405
using assms by (cases "even n", cases "0 < x",
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   406
  auto intro: DERIV_real_root[THEN DERIV_cong]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   407
              DERIV_odd_real_root[THEN DERIV_cong]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   408
              DERIV_even_real_root[THEN DERIV_cong])
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   409
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   410
subsection {* Square Root *}
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   411
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   412
definition
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   413
  sqrt :: "real \<Rightarrow> real" where
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   414
  "sqrt = root 2"
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   415
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   416
lemma pos2: "0 < (2::nat)" by simp
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   417
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   418
lemma real_sqrt_unique: "\<lbrakk>y\<twosuperior> = x; 0 \<le> y\<rbrakk> \<Longrightarrow> sqrt x = y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   419
unfolding sqrt_def by (rule real_root_pos_unique [OF pos2])
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   420
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   421
lemma real_sqrt_abs [simp]: "sqrt (x\<twosuperior>) = \<bar>x\<bar>"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   422
apply (rule real_sqrt_unique)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   423
apply (rule power2_abs)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   424
apply (rule abs_ge_zero)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   425
done
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   426
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   427
lemma real_sqrt_pow2 [simp]: "0 \<le> x \<Longrightarrow> (sqrt x)\<twosuperior> = x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   428
unfolding sqrt_def by (rule real_root_pow_pos2 [OF pos2])
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   429
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   430
lemma real_sqrt_pow2_iff [simp]: "((sqrt x)\<twosuperior> = x) = (0 \<le> x)"
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   431
apply (rule iffI)
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   432
apply (erule subst)
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   433
apply (rule zero_le_power2)
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   434
apply (erule real_sqrt_pow2)
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   435
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   436
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   437
lemma real_sqrt_zero [simp]: "sqrt 0 = 0"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   438
unfolding sqrt_def by (rule real_root_zero)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   439
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   440
lemma real_sqrt_one [simp]: "sqrt 1 = 1"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   441
unfolding sqrt_def by (rule real_root_one [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   442
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   443
lemma real_sqrt_minus: "sqrt (- x) = - sqrt x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   444
unfolding sqrt_def by (rule real_root_minus [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   445
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   446
lemma real_sqrt_mult: "sqrt (x * y) = sqrt x * sqrt y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   447
unfolding sqrt_def by (rule real_root_mult [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   448
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   449
lemma real_sqrt_inverse: "sqrt (inverse x) = inverse (sqrt x)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   450
unfolding sqrt_def by (rule real_root_inverse [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   451
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   452
lemma real_sqrt_divide: "sqrt (x / y) = sqrt x / sqrt y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   453
unfolding sqrt_def by (rule real_root_divide [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   454
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   455
lemma real_sqrt_power: "sqrt (x ^ k) = sqrt x ^ k"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   456
unfolding sqrt_def by (rule real_root_power [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   457
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   458
lemma real_sqrt_gt_zero: "0 < x \<Longrightarrow> 0 < sqrt x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   459
unfolding sqrt_def by (rule real_root_gt_zero [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   460
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   461
lemma real_sqrt_ge_zero: "0 \<le> x \<Longrightarrow> 0 \<le> sqrt x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   462
unfolding sqrt_def by (rule real_root_ge_zero [OF pos2])
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   463
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   464
lemma real_sqrt_less_mono: "x < y \<Longrightarrow> sqrt x < sqrt y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   465
unfolding sqrt_def by (rule real_root_less_mono [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   466
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   467
lemma real_sqrt_le_mono: "x \<le> y \<Longrightarrow> sqrt x \<le> sqrt y"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   468
unfolding sqrt_def by (rule real_root_le_mono [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   469
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   470
lemma real_sqrt_less_iff [simp]: "(sqrt x < sqrt y) = (x < y)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   471
unfolding sqrt_def by (rule real_root_less_iff [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   472
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   473
lemma real_sqrt_le_iff [simp]: "(sqrt x \<le> sqrt y) = (x \<le> y)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   474
unfolding sqrt_def by (rule real_root_le_iff [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   475
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   476
lemma real_sqrt_eq_iff [simp]: "(sqrt x = sqrt y) = (x = y)"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   477
unfolding sqrt_def by (rule real_root_eq_iff [OF pos2])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   478
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   479
lemmas real_sqrt_gt_0_iff [simp] = real_sqrt_less_iff [where x=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   480
lemmas real_sqrt_lt_0_iff [simp] = real_sqrt_less_iff [where y=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   481
lemmas real_sqrt_ge_0_iff [simp] = real_sqrt_le_iff [where x=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   482
lemmas real_sqrt_le_0_iff [simp] = real_sqrt_le_iff [where y=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   483
lemmas real_sqrt_eq_0_iff [simp] = real_sqrt_eq_iff [where y=0, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   484
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   485
lemmas real_sqrt_gt_1_iff [simp] = real_sqrt_less_iff [where x=1, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   486
lemmas real_sqrt_lt_1_iff [simp] = real_sqrt_less_iff [where y=1, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   487
lemmas real_sqrt_ge_1_iff [simp] = real_sqrt_le_iff [where x=1, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   488
lemmas real_sqrt_le_1_iff [simp] = real_sqrt_le_iff [where y=1, simplified]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   489
lemmas real_sqrt_eq_1_iff [simp] = real_sqrt_eq_iff [where y=1, simplified]
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   490
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   491
lemma isCont_real_sqrt: "isCont sqrt x"
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   492
unfolding sqrt_def by (rule isCont_real_root [OF pos2])
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   493
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   494
lemma DERIV_real_sqrt_generic:
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   495
  assumes "x \<noteq> 0"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   496
  assumes "x > 0 \<Longrightarrow> D = inverse (sqrt x) / 2"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   497
  assumes "x < 0 \<Longrightarrow> D = - inverse (sqrt x) / 2"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   498
  shows "DERIV sqrt x :> D"
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   499
  using assms unfolding sqrt_def
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   500
  by (auto intro!: DERIV_real_root_generic)
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   501
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   502
lemma DERIV_real_sqrt:
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   503
  "0 < x \<Longrightarrow> DERIV sqrt x :> inverse (sqrt x) / 2"
31880
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   504
  using DERIV_real_sqrt_generic by simp
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   505
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   506
declare
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   507
  DERIV_real_sqrt_generic[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
6fb86c61747c Added DERIV_intros
hoelzl
parents: 31014
diff changeset
   508
  DERIV_real_root_generic[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]
23042
492514b39956 add lemmas about continuity and derivatives of roots
huffman
parents: 23009
diff changeset
   509
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   510
lemma not_real_square_gt_zero [simp]: "(~ (0::real) < x*x) = (x = 0)"
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   511
apply auto
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   512
apply (cut_tac x = x and y = 0 in linorder_less_linear)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   513
apply (simp add: zero_less_mult_iff)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   514
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   515
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   516
lemma real_sqrt_abs2 [simp]: "sqrt(x*x) = \<bar>x\<bar>"
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   517
apply (subst power2_eq_square [symmetric])
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   518
apply (rule real_sqrt_abs)
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   519
done
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   520
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   521
lemma real_sqrt_pow2_gt_zero: "0 < x ==> 0 < (sqrt x)\<twosuperior>"
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   522
by simp (* TODO: delete *)
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   523
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   524
lemma real_sqrt_not_eq_zero: "0 < x ==> sqrt x \<noteq> 0"
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   525
by simp (* TODO: delete *)
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   526
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   527
lemma real_inv_sqrt_pow2: "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x"
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   528
by (simp add: power_inverse [symmetric])
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   529
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   530
lemma real_sqrt_eq_zero_cancel: "[| 0 \<le> x; sqrt(x) = 0|] ==> x = 0"
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   531
by simp
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   532
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   533
lemma real_sqrt_ge_one: "1 \<le> x ==> 1 \<le> sqrt x"
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   534
by simp
20687
fedb901be392 move root and sqrt stuff from Transcendental to NthRoot
huffman
parents: 20515
diff changeset
   535
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   536
lemma real_sqrt_two_gt_zero [simp]: "0 < sqrt 2"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   537
by simp
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   538
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   539
lemma real_sqrt_two_ge_zero [simp]: "0 \<le> sqrt 2"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   540
by simp
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   541
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   542
lemma real_sqrt_two_gt_one [simp]: "1 < sqrt 2"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   543
by simp
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   544
22443
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   545
lemma sqrt_divide_self_eq:
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   546
  assumes nneg: "0 \<le> x"
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   547
  shows "sqrt x / x = inverse (sqrt x)"
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   548
proof cases
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   549
  assume "x=0" thus ?thesis by simp
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   550
next
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   551
  assume nz: "x\<noteq>0" 
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   552
  hence pos: "0<x" using nneg by arith
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   553
  show ?thesis
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   554
  proof (rule right_inverse_eq [THEN iffD1, THEN sym]) 
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   555
    show "sqrt x / x \<noteq> 0" by (simp add: divide_inverse nneg nz) 
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   556
    show "inverse (sqrt x) / (sqrt x / x) = 1"
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   557
      by (simp add: divide_inverse mult_assoc [symmetric] 
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   558
                  power2_eq_square [symmetric] real_inv_sqrt_pow2 pos nz) 
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   559
  qed
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   560
qed
346729a55460 move sqrt_divide_self_eq to NthRoot.thy
huffman
parents: 21865
diff changeset
   561
22721
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   562
lemma real_divide_square_eq [simp]: "(((r::real) * a) / (r * r)) = a / r"
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   563
apply (simp add: divide_inverse)
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   564
apply (case_tac "r=0")
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   565
apply (auto simp add: mult_ac)
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   566
done
d9be18bd7a28 moved root and sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 22630
diff changeset
   567
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   568
lemma lemma_real_divide_sqrt_less: "0 < u ==> u / sqrt 2 < u"
35216
7641e8d831d2 get rid of many duplicate simp rule warnings
huffman
parents: 31880
diff changeset
   569
by (simp add: divide_less_eq)
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   570
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   571
lemma four_x_squared: 
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   572
  fixes x::real
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   573
  shows "4 * x\<twosuperior> = (2 * x)\<twosuperior>"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   574
by (simp add: power2_eq_square)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   575
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   576
subsection {* Square Root of Sum of Squares *}
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   577
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   578
lemma real_sqrt_mult_self_sum_ge_zero [simp]: "0 \<le> sqrt(x*x + y*y)"
22968
huffman
parents: 22961
diff changeset
   579
by (rule real_sqrt_ge_zero [OF sum_squares_ge_zero])
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   580
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   581
lemma real_sqrt_sum_squares_ge_zero [simp]: "0 \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
22961
e499ded5d0fc remove redundant lemmas
huffman
parents: 22956
diff changeset
   582
by simp
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   583
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   584
declare real_sqrt_sum_squares_ge_zero [THEN abs_of_nonneg, simp]
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   585
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   586
lemma real_sqrt_sum_squares_mult_ge_zero [simp]:
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   587
     "0 \<le> sqrt ((x\<twosuperior> + y\<twosuperior>)*(xa\<twosuperior> + ya\<twosuperior>))"
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   588
by (auto intro!: real_sqrt_ge_zero simp add: zero_le_mult_iff)
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   589
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   590
lemma real_sqrt_sum_squares_mult_squared_eq [simp]:
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   591
     "sqrt ((x\<twosuperior> + y\<twosuperior>) * (xa\<twosuperior> + ya\<twosuperior>)) ^ 2 = (x\<twosuperior> + y\<twosuperior>) * (xa\<twosuperior> + ya\<twosuperior>)"
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   592
by (auto simp add: zero_le_mult_iff)
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   593
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   594
lemma real_sqrt_sum_squares_eq_cancel: "sqrt (x\<twosuperior> + y\<twosuperior>) = x \<Longrightarrow> y = 0"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   595
by (drule_tac f = "%x. x\<twosuperior>" in arg_cong, simp)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   596
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   597
lemma real_sqrt_sum_squares_eq_cancel2: "sqrt (x\<twosuperior> + y\<twosuperior>) = y \<Longrightarrow> x = 0"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   598
by (drule_tac f = "%x. x\<twosuperior>" in arg_cong, simp)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   599
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   600
lemma real_sqrt_sum_squares_ge1 [simp]: "x \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   601
by (rule power2_le_imp_le, simp_all)
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   602
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   603
lemma real_sqrt_sum_squares_ge2 [simp]: "y \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   604
by (rule power2_le_imp_le, simp_all)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   605
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   606
lemma real_sqrt_ge_abs1 [simp]: "\<bar>x\<bar> \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
22856
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   607
by (rule power2_le_imp_le, simp_all)
eb0e0582092a cleaned up
huffman
parents: 22721
diff changeset
   608
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   609
lemma real_sqrt_ge_abs2 [simp]: "\<bar>y\<bar> \<le> sqrt (x\<twosuperior> + y\<twosuperior>)"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   610
by (rule power2_le_imp_le, simp_all)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   611
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   612
lemma le_real_sqrt_sumsq [simp]: "x \<le> sqrt (x * x + y * y)"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   613
by (simp add: power2_eq_square [symmetric])
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   614
22858
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   615
lemma real_sqrt_sum_squares_triangle_ineq:
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   616
  "sqrt ((a + c)\<twosuperior> + (b + d)\<twosuperior>) \<le> sqrt (a\<twosuperior> + b\<twosuperior>) + sqrt (c\<twosuperior> + d\<twosuperior>)"
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   617
apply (rule power2_le_imp_le, simp)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   618
apply (simp add: power2_sum)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   619
apply (simp only: mult_assoc right_distrib [symmetric])
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   620
apply (rule mult_left_mono)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   621
apply (rule power2_le_imp_le)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   622
apply (simp add: power2_sum power_mult_distrib)
23477
f4b83f03cac9 tuned and renamed group_eq_simps and ring_eq_simps
nipkow
parents: 23475
diff changeset
   623
apply (simp add: ring_distribs)
22858
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   624
apply (subgoal_tac "0 \<le> b\<twosuperior> * c\<twosuperior> + a\<twosuperior> * d\<twosuperior> - 2 * (a * c) * (b * d)", simp)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   625
apply (rule_tac b="(a * d - b * c)\<twosuperior>" in ord_le_eq_trans)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   626
apply (rule zero_le_power2)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   627
apply (simp add: power2_diff power_mult_distrib)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   628
apply (simp add: mult_nonneg_nonneg)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   629
apply simp
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   630
apply (simp add: add_increasing)
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   631
done
5ca5d1cce388 add lemma real_sqrt_sum_squares_triangle_ineq
huffman
parents: 22856
diff changeset
   632
23122
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   633
lemma real_sqrt_sum_squares_less:
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   634
  "\<lbrakk>\<bar>x\<bar> < u / sqrt 2; \<bar>y\<bar> < u / sqrt 2\<rbrakk> \<Longrightarrow> sqrt (x\<twosuperior> + y\<twosuperior>) < u"
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   635
apply (rule power2_less_imp_less, simp)
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   636
apply (drule power_strict_mono [OF _ abs_ge_zero pos2])
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   637
apply (drule power_strict_mono [OF _ abs_ge_zero pos2])
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   638
apply (simp add: power_divide)
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   639
apply (drule order_le_less_trans [OF abs_ge_zero])
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   640
apply (simp add: zero_less_divide_iff)
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   641
done
3d853d6f2f7d add lemma real_sqrt_sum_squares_less
huffman
parents: 23069
diff changeset
   642
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   643
text{*Needed for the infinitely close relation over the nonstandard
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   644
    complex numbers*}
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   645
lemma lemma_sqrt_hcomplex_capprox:
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   646
     "[| 0 < u; x < u/2; y < u/2; 0 \<le> x; 0 \<le> y |] ==> sqrt (x\<twosuperior> + y\<twosuperior>) < u"
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   647
apply (rule_tac y = "u/sqrt 2" in order_le_less_trans)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   648
apply (erule_tac [2] lemma_real_divide_sqrt_less)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   649
apply (rule power2_le_imp_le)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   650
apply (auto simp add: real_0_le_divide_iff power_divide)
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   651
apply (rule_tac t = "u\<twosuperior>" in real_sum_of_halves [THEN subst])
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   652
apply (rule add_mono)
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 28952
diff changeset
   653
apply (auto simp add: four_x_squared intro: power_mono)
23049
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   654
done
11607c283074 moved sqrt lemmas from Transcendental.thy to NthRoot.thy
huffman
parents: 23047
diff changeset
   655
22956
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   656
text "Legacy theorem names:"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   657
lemmas real_root_pos2 = real_root_power_cancel
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   658
lemmas real_root_pos_pos = real_root_gt_zero [THEN order_less_imp_le]
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   659
lemmas real_root_pos_pos_le = real_root_ge_zero
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   660
lemmas real_sqrt_mult_distrib = real_sqrt_mult
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   661
lemmas real_sqrt_mult_distrib2 = real_sqrt_mult
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   662
lemmas real_sqrt_eq_zero_cancel_iff = real_sqrt_eq_0_iff
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   663
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   664
(* needed for CauchysMeanTheorem.het_base from AFP *)
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   665
lemma real_root_pos: "0 < x \<Longrightarrow> root (Suc n) (x ^ (Suc n)) = x"
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   666
by (rule real_root_power_cancel [OF zero_less_Suc order_less_imp_le])
617140080e6a define roots of negative reals so that many lemmas no longer require side conditions; simplification solves more goals than previously
huffman
parents: 22943
diff changeset
   667
14324
c9c6832f9b22 converting Hyperreal/NthRoot to Isar
paulson
parents: 14268
diff changeset
   668
end