src/HOL/Computational_Algebra/Formal_Power_Series.thy
author haftmann
Thu, 22 Nov 2018 10:06:31 +0000
changeset 69325 4b6ddc5989fc
parent 69260 0a9688695a1b
child 69597 ff784d5a5bfb
permissions -rw-r--r--
removed legacy input syntax
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
65435
378175f44328 tuned headers;
wenzelm
parents: 65417
diff changeset
     1
(*  Title:      HOL/Computational_Algebra/Formal_Power_Series.thy
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
     2
    Author:     Amine Chaieb, University of Cambridge
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
     3
*)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
     4
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
     5
section \<open>A formalization of formal power series\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
     6
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
     7
theory Formal_Power_Series
65417
fc41a5650fb1 session containing computational algebra
haftmann
parents: 65398
diff changeset
     8
imports
fc41a5650fb1 session containing computational algebra
haftmann
parents: 65398
diff changeset
     9
  Complex_Main
fc41a5650fb1 session containing computational algebra
haftmann
parents: 65398
diff changeset
    10
  Euclidean_Algorithm
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    11
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    12
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    13
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    14
subsection \<open>The type of formal power series\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    15
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 48757
diff changeset
    16
typedef 'a fps = "{f :: nat \<Rightarrow> 'a. True}"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    17
  morphisms fps_nth Abs_fps
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    18
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    19
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    20
notation fps_nth (infixl "$" 75)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    21
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    22
lemma expand_fps_eq: "p = q \<longleftrightarrow> (\<forall>n. p $ n = q $ n)"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
    23
  by (simp add: fps_nth_inject [symmetric] fun_eq_iff)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    24
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    25
lemma fps_ext: "(\<And>n. p $ n = q $ n) \<Longrightarrow> p = q"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    26
  by (simp add: expand_fps_eq)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    27
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    28
lemma fps_nth_Abs_fps [simp]: "Abs_fps f $ n = f n"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    29
  by (simp add: Abs_fps_inverse)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    30
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    31
text \<open>Definition of the basic elements 0 and 1 and the basic operations of addition,
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    32
  negation and multiplication.\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    33
36409
d323e7773aa8 use new classes (linordered_)field_inverse_zero
haftmann
parents: 36350
diff changeset
    34
instantiation fps :: (zero) zero
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    35
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    36
  definition fps_zero_def: "0 = Abs_fps (\<lambda>n. 0)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    37
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    38
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    39
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    40
lemma fps_zero_nth [simp]: "0 $ n = 0"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    41
  unfolding fps_zero_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    42
36409
d323e7773aa8 use new classes (linordered_)field_inverse_zero
haftmann
parents: 36350
diff changeset
    43
instantiation fps :: ("{one, zero}") one
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    44
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    45
  definition fps_one_def: "1 = Abs_fps (\<lambda>n. if n = 0 then 1 else 0)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    46
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    47
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    48
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
    49
lemma fps_one_nth [simp]: "1 $ n = (if n = 0 then 1 else 0)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    50
  unfolding fps_one_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    51
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
    52
instantiation fps :: (plus) plus
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    53
begin
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66817
diff changeset
    54
  definition fps_plus_def: "(+) = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n + g $ n))"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    55
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    56
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    57
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    58
lemma fps_add_nth [simp]: "(f + g) $ n = f $ n + g $ n"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    59
  unfolding fps_plus_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    60
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    61
instantiation fps :: (minus) minus
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    62
begin
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66817
diff changeset
    63
  definition fps_minus_def: "(-) = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n - g $ n))"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    64
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    65
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    66
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    67
lemma fps_sub_nth [simp]: "(f - g) $ n = f $ n - g $ n"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    68
  unfolding fps_minus_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    69
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    70
instantiation fps :: (uminus) uminus
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    71
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    72
  definition fps_uminus_def: "uminus = (\<lambda>f. Abs_fps (\<lambda>n. - (f $ n)))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    73
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    74
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    75
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    76
lemma fps_neg_nth [simp]: "(- f) $ n = - (f $ n)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    77
  unfolding fps_uminus_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    78
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
    79
instantiation fps :: ("{comm_monoid_add, times}") times
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    80
begin
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68975
diff changeset
    81
  definition fps_times_def: "(*) = (\<lambda>f g. Abs_fps (\<lambda>n. \<Sum>i=0..n. f $ i * g $ (n - i)))"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    82
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    83
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    84
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    85
lemma fps_mult_nth: "(f * g) $ n = (\<Sum>i=0..n. f$i * g$(n - i))"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    86
  unfolding fps_times_def by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    87
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
    88
lemma fps_mult_nth_0 [simp]: "(f * g) $ 0 = f $ 0 * g $ 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
    89
  unfolding fps_times_def by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
    90
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
    91
declare atLeastAtMost_iff [presburger]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
    92
declare Bex_def [presburger]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
    93
declare Ball_def [presburger]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    94
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    95
lemma mult_delta_left:
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    96
  fixes x y :: "'a::mult_zero"
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    97
  shows "(if b then x else 0) * y = (if b then x * y else 0)"
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    98
  by simp
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    99
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   100
lemma mult_delta_right:
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   101
  fixes x y :: "'a::mult_zero"
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   102
  shows "x * (if b then y else 0) = (if b then x * y else 0)"
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   103
  by simp
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   104
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   105
lemma cond_value_iff: "f (if b then x else y) = (if b then f x else f y)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   106
  by auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   107
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   108
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   109
subsection \<open>Formal power series form a commutative ring with unity, if the range of sequences
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   110
  they represent is a commutative ring with unity\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   111
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   112
instance fps :: (semigroup_add) semigroup_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   113
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   114
  fix a b c :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   115
  show "a + b + c = a + (b + c)"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   116
    by (simp add: fps_ext add.assoc)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   117
qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   118
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   119
instance fps :: (ab_semigroup_add) ab_semigroup_add
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   120
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   121
  fix a b :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   122
  show "a + b = b + a"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   123
    by (simp add: fps_ext add.commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   124
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   125
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   126
lemma fps_mult_assoc_lemma:
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   127
  fixes k :: nat
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   128
    and f :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   129
  shows "(\<Sum>j=0..k. \<Sum>i=0..j. f i (j - i) (n - j)) =
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   130
         (\<Sum>j=0..k. \<Sum>i=0..k - j. f j i (n - j - i))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   131
  by (induct k) (simp_all add: Suc_diff_le sum.distrib add.assoc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   132
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   133
instance fps :: (semiring_0) semigroup_mult
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   134
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   135
  fix a b c :: "'a fps"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   136
  show "(a * b) * c = a * (b * c)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   137
  proof (rule fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   138
    fix n :: nat
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   139
    have "(\<Sum>j=0..n. \<Sum>i=0..j. a$i * b$(j - i) * c$(n - j)) =
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   140
          (\<Sum>j=0..n. \<Sum>i=0..n - j. a$j * b$i * c$(n - j - i))"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   141
      by (rule fps_mult_assoc_lemma)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   142
    then show "((a * b) * c) $ n = (a * (b * c)) $ n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   143
      by (simp add: fps_mult_nth sum_distrib_left sum_distrib_right mult.assoc)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   144
  qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   145
qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   146
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   147
lemma fps_mult_commute_lemma:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   148
  fixes n :: nat
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   149
    and f :: "nat \<Rightarrow> nat \<Rightarrow> 'a::comm_monoid_add"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   150
  shows "(\<Sum>i=0..n. f i (n - i)) = (\<Sum>i=0..n. f (n - i) i)"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66817
diff changeset
   151
  by (rule sum.reindex_bij_witness[where i="(-) n" and j="(-) n"]) auto
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   152
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   153
instance fps :: (comm_semiring_0) ab_semigroup_mult
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   154
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   155
  fix a b :: "'a fps"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   156
  show "a * b = b * a"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   157
  proof (rule fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   158
    fix n :: nat
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   159
    have "(\<Sum>i=0..n. a$i * b$(n - i)) = (\<Sum>i=0..n. a$(n - i) * b$i)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   160
      by (rule fps_mult_commute_lemma)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   161
    then show "(a * b) $ n = (b * a) $ n"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   162
      by (simp add: fps_mult_nth mult.commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   163
  qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   164
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   165
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   166
instance fps :: (monoid_add) monoid_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   167
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   168
  fix a :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   169
  show "0 + a = a" by (simp add: fps_ext)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   170
  show "a + 0 = a" by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   171
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   172
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   173
instance fps :: (comm_monoid_add) comm_monoid_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   174
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   175
  fix a :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   176
  show "0 + a = a" by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   177
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   178
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   179
instance fps :: (semiring_1) monoid_mult
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   180
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   181
  fix a :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   182
  show "1 * a = a"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   183
    by (simp add: fps_ext fps_mult_nth mult_delta_left sum.delta)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   184
  show "a * 1 = a"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   185
    by (simp add: fps_ext fps_mult_nth mult_delta_right sum.delta')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   186
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   187
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   188
instance fps :: (cancel_semigroup_add) cancel_semigroup_add
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   189
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   190
  fix a b c :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   191
  show "b = c" if "a + b = a + c"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   192
    using that by (simp add: expand_fps_eq)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   193
  show "b = c" if "b + a = c + a"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   194
    using that by (simp add: expand_fps_eq)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   195
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   196
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   197
instance fps :: (cancel_ab_semigroup_add) cancel_ab_semigroup_add
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   198
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   199
  fix a b c :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   200
  show "a + b - a = b"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   201
    by (simp add: expand_fps_eq)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   202
  show "a - b - c = a - (b + c)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   203
    by (simp add: expand_fps_eq diff_diff_eq)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   204
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   205
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   206
instance fps :: (cancel_comm_monoid_add) cancel_comm_monoid_add ..
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   207
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   208
instance fps :: (group_add) group_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   209
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   210
  fix a b :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   211
  show "- a + a = 0" by (simp add: fps_ext)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
   212
  show "a + - b = a - b" by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   213
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   214
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   215
instance fps :: (ab_group_add) ab_group_add
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   216
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   217
  fix a b :: "'a fps"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   218
  show "- a + a = 0" by (simp add: fps_ext)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   219
  show "a - b = a + - b" by (simp add: fps_ext)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   220
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   221
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   222
instance fps :: (zero_neq_one) zero_neq_one
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60567
diff changeset
   223
  by standard (simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   224
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   225
instance fps :: (semiring_0) semiring
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   226
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   227
  fix a b c :: "'a fps"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   228
  show "(a + b) * c = a * c + b * c"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   229
    by (simp add: expand_fps_eq fps_mult_nth distrib_right sum.distrib)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   230
  show "a * (b + c) = a * b + a * c"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   231
    by (simp add: expand_fps_eq fps_mult_nth distrib_left sum.distrib)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   232
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   233
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   234
instance fps :: (semiring_0) semiring_0
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   235
proof
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   236
  fix a :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   237
  show "0 * a = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   238
    by (simp add: fps_ext fps_mult_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   239
  show "a * 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   240
    by (simp add: fps_ext fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   241
qed
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   242
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   243
instance fps :: (semiring_0_cancel) semiring_0_cancel ..
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   244
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   245
instance fps :: (semiring_1) semiring_1 ..
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   246
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   247
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   248
subsection \<open>Selection of the nth power of the implicit variable in the infinite sum\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   249
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   250
lemma fps_square_nth: "(f^2) $ n = (\<Sum>k\<le>n. f $ k * f $ (n - k))"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   251
  by (simp add: power2_eq_square fps_mult_nth atLeast0AtMost)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   252
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   253
lemma fps_nonzero_nth: "f \<noteq> 0 \<longleftrightarrow> (\<exists> n. f $n \<noteq> 0)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   254
  by (simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   255
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   256
lemma fps_nonzero_nth_minimal: "f \<noteq> 0 \<longleftrightarrow> (\<exists>n. f $ n \<noteq> 0 \<and> (\<forall>m < n. f $ m = 0))"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   257
  (is "?lhs \<longleftrightarrow> ?rhs")
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   258
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   259
  let ?n = "LEAST n. f $ n \<noteq> 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   260
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   261
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   262
    from that have "\<exists>n. f $ n \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   263
      by (simp add: fps_nonzero_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   264
    then have "f $ ?n \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   265
      by (rule LeastI_ex)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   266
    moreover have "\<forall>m<?n. f $ m = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   267
      by (auto dest: not_less_Least)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   268
    ultimately have "f $ ?n \<noteq> 0 \<and> (\<forall>m<?n. f $ m = 0)" ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   269
    then show ?thesis ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   270
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   271
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   272
    using that by (auto simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   273
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   274
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   275
lemma fps_eq_iff: "f = g \<longleftrightarrow> (\<forall>n. f $ n = g $n)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   276
  by (rule expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   277
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   278
lemma fps_sum_nth: "sum f S $ n = sum (\<lambda>k. (f k) $ n) S"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   279
proof (cases "finite S")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   280
  case True
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   281
  then show ?thesis by (induct set: finite) auto
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   282
next
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   283
  case False
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   284
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   285
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   286
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   287
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   288
subsection \<open>Injection of the basic ring elements and multiplication by scalars\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   289
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   290
definition "fps_const c = Abs_fps (\<lambda>n. if n = 0 then c else 0)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   291
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   292
lemma fps_nth_fps_const [simp]: "fps_const c $ n = (if n = 0 then c else 0)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   293
  unfolding fps_const_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   294
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   295
lemma fps_const_0_eq_0 [simp]: "fps_const 0 = 0"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   296
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   297
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   298
lemma fps_const_1_eq_1 [simp]: "fps_const 1 = 1"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   299
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   300
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   301
lemma fps_const_neg [simp]: "- (fps_const (c::'a::ring)) = fps_const (- c)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   302
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   303
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   304
lemma fps_const_add [simp]: "fps_const (c::'a::monoid_add) + fps_const d = fps_const (c + d)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   305
  by (simp add: fps_ext)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   306
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   307
lemma fps_const_sub [simp]: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
   308
  by (simp add: fps_ext)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   309
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   310
lemma fps_const_mult[simp]: "fps_const (c::'a::ring) * fps_const d = fps_const (c * d)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   311
  by (simp add: fps_eq_iff fps_mult_nth sum.neutral)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   312
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   313
lemma fps_const_add_left: "fps_const (c::'a::monoid_add) + f =
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   314
    Abs_fps (\<lambda>n. if n = 0 then c + f$0 else f$n)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   315
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   316
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   317
lemma fps_const_add_right: "f + fps_const (c::'a::monoid_add) =
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   318
    Abs_fps (\<lambda>n. if n = 0 then f$0 + c else f$n)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   319
  by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   320
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   321
lemma fps_const_mult_left: "fps_const (c::'a::semiring_0) * f = Abs_fps (\<lambda>n. c * f$n)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   322
  unfolding fps_eq_iff fps_mult_nth
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   323
  by (simp add: fps_const_def mult_delta_left sum.delta)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   324
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   325
lemma fps_const_mult_right: "f * fps_const (c::'a::semiring_0) = Abs_fps (\<lambda>n. f$n * c)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   326
  unfolding fps_eq_iff fps_mult_nth
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   327
  by (simp add: fps_const_def mult_delta_right sum.delta')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   328
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   329
lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::semiring_1) * f)$n = c* f$n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   330
  by (simp add: fps_mult_nth mult_delta_left sum.delta)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   331
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   332
lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   333
  by (simp add: fps_mult_nth mult_delta_right sum.delta')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   334
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   335
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   336
subsection \<open>Formal power series form an integral domain\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   337
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   338
instance fps :: (ring) ring ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   339
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   340
instance fps :: (ring_1) ring_1
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
   341
  by (intro_classes, auto simp add: distrib_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   342
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   343
instance fps :: (comm_ring_1) comm_ring_1
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
   344
  by (intro_classes, auto simp add: distrib_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   345
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   346
instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   347
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   348
  fix a b :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   349
  assume "a \<noteq> 0" and "b \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   350
  then obtain i j where i: "a $ i \<noteq> 0" "\<forall>k<i. a $ k = 0" and j: "b $ j \<noteq> 0" "\<forall>k<j. b $ k =0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   351
    unfolding fps_nonzero_nth_minimal
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   352
    by blast+
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   353
  have "(a * b) $ (i + j) = (\<Sum>k=0..i+j. a $ k * b $ (i + j - k))"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   354
    by (rule fps_mult_nth)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   355
  also have "\<dots> = (a $ i * b $ (i + j - i)) + (\<Sum>k\<in>{0..i+j} - {i}. a $ k * b $ (i + j - k))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   356
    by (rule sum.remove) simp_all
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   357
  also have "(\<Sum>k\<in>{0..i+j}-{i}. a $ k * b $ (i + j - k)) = 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   358
  proof (rule sum.neutral [rule_format])
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   359
    fix k assume "k \<in> {0..i+j} - {i}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   360
    then have "k < i \<or> i+j-k < j"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   361
      by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   362
    then show "a $ k * b $ (i + j - k) = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   363
      using i j by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   364
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   365
  also have "a $ i * b $ (i + j - i) + 0 = a $ i * b $ j"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   366
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   367
  also have "a $ i * b $ j \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   368
    using i j by simp
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   369
  finally have "(a*b) $ (i+j) \<noteq> 0" .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   370
  then show "a * b \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   371
    unfolding fps_nonzero_nth by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   372
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   373
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   374
instance fps :: (ring_1_no_zero_divisors) ring_1_no_zero_divisors ..
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   375
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   376
instance fps :: (idom) idom ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   377
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   378
lemma numeral_fps_const: "numeral k = fps_const (numeral k)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   379
  by (induct k) (simp_all only: numeral.simps fps_const_1_eq_1
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   380
    fps_const_add [symmetric])
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   381
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   382
lemma neg_numeral_fps_const:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   383
  "(- numeral k :: 'a :: ring_1 fps) = fps_const (- numeral k)"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   384
  by (simp add: numeral_fps_const)
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   385
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   386
lemma fps_numeral_nth: "numeral n $ i = (if i = 0 then numeral n else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   387
  by (simp add: numeral_fps_const)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   388
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   389
lemma fps_numeral_nth_0 [simp]: "numeral n $ 0 = numeral n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   390
  by (simp add: numeral_fps_const)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   391
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   392
lemma fps_of_nat: "fps_const (of_nat c) = of_nat c"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   393
  by (induction c) (simp_all add: fps_const_add [symmetric] del: fps_const_add)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   394
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
   395
lemma numeral_neq_fps_zero [simp]: "(numeral f :: 'a :: field_char_0 fps) \<noteq> 0"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
   396
proof
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
   397
  assume "numeral f = (0 :: 'a fps)"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
   398
  from arg_cong[of _ _ "\<lambda>F. F $ 0", OF this] show False by simp
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
   399
qed 
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   400
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   401
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   402
subsection \<open>The efps_Xtractor series fps_X\<close>
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   403
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   404
lemma minus_one_power_iff: "(- (1::'a::comm_ring_1)) ^ n = (if even n then 1 else - 1)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   405
  by (induct n) auto
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   406
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   407
definition "fps_X = Abs_fps (\<lambda>n. if n = 1 then 1 else 0)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   408
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   409
lemma fps_X_mult_nth [simp]:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   410
  "(fps_X * (f :: 'a::semiring_1 fps)) $n = (if n = 0 then 0 else f $ (n - 1))"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   411
proof (cases "n = 0")
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   412
  case False
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   413
  have "(fps_X * f) $n = (\<Sum>i = 0..n. fps_X $ i * f $ (n - i))"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   414
    by (simp add: fps_mult_nth)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   415
  also have "\<dots> = f $ (n - 1)"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   416
    using False by (simp add: fps_X_def mult_delta_left sum.delta)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   417
  finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   418
    using False by simp
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   419
next
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   420
  case True
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   421
  then show ?thesis
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   422
    by (simp add: fps_mult_nth fps_X_def)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   423
qed
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   424
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   425
lemma fps_X_mult_right_nth[simp]:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   426
  "((a::'a::semiring_1 fps) * fps_X) $ n = (if n = 0 then 0 else a $ (n - 1))"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   427
proof -
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   428
  have "(a * fps_X) $ n = (\<Sum>i = 0..n. a $ i * (if n - i = Suc 0 then 1 else 0))"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   429
    by (simp add: fps_times_def fps_X_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   430
  also have "\<dots> = (\<Sum>i = 0..n. if i = n - 1 then if n = 0 then 0 else a $ i else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   431
    by (intro sum.cong) auto
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   432
  also have "\<dots> = (if n = 0 then 0 else a $ (n - 1))" by (simp add: sum.delta)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   433
  finally show ?thesis .
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   434
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   435
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   436
lemma fps_mult_fps_X_commute: "fps_X * (a :: 'a :: semiring_1 fps) = a * fps_X" 
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
   437
  by (simp add: fps_eq_iff)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   438
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   439
lemma fps_X_power_iff: "fps_X ^ n = Abs_fps (\<lambda>m. if m = n then 1 else 0)"
66466
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
   440
  by (induction n) (auto simp: fps_eq_iff)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   441
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   442
lemma fps_X_nth[simp]: "fps_X$n = (if n = 1 then 1 else 0)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   443
  by (simp add: fps_X_def)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   444
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   445
lemma fps_X_power_nth[simp]: "(fps_X^k) $n = (if n = k then 1 else 0::'a::comm_ring_1)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   446
  by (simp add: fps_X_power_iff)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   447
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   448
lemma fps_X_power_mult_nth: "(fps_X^k * (f :: 'a::comm_ring_1 fps)) $n = (if n < k then 0 else f $ (n - k))"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   449
  apply (induct k arbitrary: n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   450
  apply simp
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   451
  unfolding power_Suc mult.assoc
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   452
  apply (case_tac n)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   453
  apply auto
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   454
  done
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   455
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   456
lemma fps_X_power_mult_right_nth:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   457
    "((f :: 'a::comm_ring_1 fps) * fps_X^k) $n = (if n < k then 0 else f $ (n - k))"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   458
  by (metis fps_X_power_mult_nth mult.commute)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   459
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   460
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   461
lemma fps_X_neq_fps_const [simp]: "(fps_X :: 'a :: zero_neq_one fps) \<noteq> fps_const c"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   462
proof
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   463
  assume "(fps_X::'a fps) = fps_const (c::'a)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   464
  hence "fps_X$1 = (fps_const (c::'a))$1" by (simp only:)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   465
  thus False by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   466
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   467
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   468
lemma fps_X_neq_zero [simp]: "(fps_X :: 'a :: zero_neq_one fps) \<noteq> 0"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   469
  by (simp only: fps_const_0_eq_0[symmetric] fps_X_neq_fps_const) simp
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   470
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   471
lemma fps_X_neq_one [simp]: "(fps_X :: 'a :: zero_neq_one fps) \<noteq> 1"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   472
  by (simp only: fps_const_1_eq_1[symmetric] fps_X_neq_fps_const) simp
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   473
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   474
lemma fps_X_neq_numeral [simp]: "(fps_X :: 'a :: {semiring_1,zero_neq_one} fps) \<noteq> numeral c"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   475
  by (simp only: numeral_fps_const fps_X_neq_fps_const) simp
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   476
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   477
lemma fps_X_pow_eq_fps_X_pow_iff [simp]:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   478
  "(fps_X :: ('a :: {comm_ring_1}) fps) ^ m = fps_X ^ n \<longleftrightarrow> m = n"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   479
proof
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   480
  assume "(fps_X :: 'a fps) ^ m = fps_X ^ n"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   481
  hence "(fps_X :: 'a fps) ^ m $ m = fps_X ^ n $ m" by (simp only:)
62390
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
   482
  thus "m = n" by (simp split: if_split_asm)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   483
qed simp_all
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   484
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   485
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   486
subsection \<open>Subdegrees\<close>
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   487
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   488
definition subdegree :: "('a::zero) fps \<Rightarrow> nat" where
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   489
  "subdegree f = (if f = 0 then 0 else LEAST n. f$n \<noteq> 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   490
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   491
lemma subdegreeI:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   492
  assumes "f $ d \<noteq> 0" and "\<And>i. i < d \<Longrightarrow> f $ i = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   493
  shows   "subdegree f = d"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   494
proof-
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   495
  from assms(1) have "f \<noteq> 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   496
  moreover from assms(1) have "(LEAST i. f $ i \<noteq> 0) = d"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   497
  proof (rule Least_equality)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   498
    fix e assume "f $ e \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   499
    with assms(2) have "\<not>(e < d)" by blast
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   500
    thus "e \<ge> d" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   501
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   502
  ultimately show ?thesis unfolding subdegree_def by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   503
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   504
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   505
lemma nth_subdegree_nonzero [simp,intro]: "f \<noteq> 0 \<Longrightarrow> f $ subdegree f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   506
proof-
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   507
  assume "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   508
  hence "subdegree f = (LEAST n. f $ n \<noteq> 0)" by (simp add: subdegree_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   509
  also from \<open>f \<noteq> 0\<close> have "\<exists>n. f$n \<noteq> 0" using fps_nonzero_nth by blast
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   510
  from LeastI_ex[OF this] have "f $ (LEAST n. f $ n \<noteq> 0) \<noteq> 0" .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   511
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   512
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   513
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   514
lemma nth_less_subdegree_zero [dest]: "n < subdegree f \<Longrightarrow> f $ n = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   515
proof (cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   516
  assume "f \<noteq> 0" and less: "n < subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   517
  note less
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   518
  also from \<open>f \<noteq> 0\<close> have "subdegree f = (LEAST n. f $ n \<noteq> 0)" by (simp add: subdegree_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   519
  finally show "f $ n = 0" using not_less_Least by blast
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   520
qed simp_all
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   521
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   522
lemma subdegree_geI:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   523
  assumes "f \<noteq> 0" "\<And>i. i < n \<Longrightarrow> f$i = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   524
  shows   "subdegree f \<ge> n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   525
proof (rule ccontr)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   526
  assume "\<not>(subdegree f \<ge> n)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   527
  with assms(2) have "f $ subdegree f = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   528
  moreover from assms(1) have "f $ subdegree f \<noteq> 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   529
  ultimately show False by contradiction
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   530
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   531
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   532
lemma subdegree_greaterI:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   533
  assumes "f \<noteq> 0" "\<And>i. i \<le> n \<Longrightarrow> f$i = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   534
  shows   "subdegree f > n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   535
proof (rule ccontr)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   536
  assume "\<not>(subdegree f > n)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   537
  with assms(2) have "f $ subdegree f = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   538
  moreover from assms(1) have "f $ subdegree f \<noteq> 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   539
  ultimately show False by contradiction
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   540
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   541
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   542
lemma subdegree_leI:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   543
  "f $ n \<noteq> 0 \<Longrightarrow> subdegree f \<le> n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   544
  by (rule leI) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   545
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   546
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   547
lemma subdegree_0 [simp]: "subdegree 0 = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   548
  by (simp add: subdegree_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   549
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   550
lemma subdegree_1 [simp]: "subdegree (1 :: ('a :: zero_neq_one) fps) = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   551
  by (auto intro!: subdegreeI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   552
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   553
lemma subdegree_fps_X [simp]: "subdegree (fps_X :: ('a :: zero_neq_one) fps) = 1"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   554
  by (auto intro!: subdegreeI simp: fps_X_def)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   555
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   556
lemma subdegree_fps_const [simp]: "subdegree (fps_const c) = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   557
  by (cases "c = 0") (auto intro!: subdegreeI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   558
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   559
lemma subdegree_numeral [simp]: "subdegree (numeral n) = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   560
  by (simp add: numeral_fps_const)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   561
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   562
lemma subdegree_eq_0_iff: "subdegree f = 0 \<longleftrightarrow> f = 0 \<or> f $ 0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   563
proof (cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   564
  assume "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   565
  thus ?thesis
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   566
    using nth_subdegree_nonzero[OF \<open>f \<noteq> 0\<close>] by (fastforce intro!: subdegreeI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   567
qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   568
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   569
lemma subdegree_eq_0 [simp]: "f $ 0 \<noteq> 0 \<Longrightarrow> subdegree f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   570
  by (simp add: subdegree_eq_0_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   571
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   572
lemma nth_subdegree_mult [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   573
  fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   574
  shows "(f * g) $ (subdegree f + subdegree g) = f $ subdegree f * g $ subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   575
proof-
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   576
  let ?n = "subdegree f + subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   577
  have "(f * g) $ ?n = (\<Sum>i=0..?n. f$i * g$(?n-i))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   578
    by (simp add: fps_mult_nth)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   579
  also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   580
  proof (intro sum.cong)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   581
    fix x assume x: "x \<in> {0..?n}"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   582
    hence "x = subdegree f \<or> x < subdegree f \<or> ?n - x < subdegree g" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   583
    thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   584
      by (elim disjE conjE) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   585
  qed auto
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   586
  also have "... = f $ subdegree f * g $ subdegree g" by (simp add: sum.delta)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   587
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   588
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   589
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   590
lemma subdegree_mult [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   591
  assumes "f \<noteq> 0" "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   592
  shows "subdegree ((f :: ('a :: {ring_no_zero_divisors}) fps) * g) = subdegree f + subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   593
proof (rule subdegreeI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   594
  let ?n = "subdegree f + subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   595
  have "(f * g) $ ?n = (\<Sum>i=0..?n. f$i * g$(?n-i))" by (simp add: fps_mult_nth)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   596
  also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   597
  proof (intro sum.cong)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   598
    fix x assume x: "x \<in> {0..?n}"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   599
    hence "x = subdegree f \<or> x < subdegree f \<or> ?n - x < subdegree g" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   600
    thus "f $ x * g $ (?n - x) = (if x = subdegree f then f $ x * g $ (?n - x) else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   601
      by (elim disjE conjE) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   602
  qed auto
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   603
  also have "... = f $ subdegree f * g $ subdegree g" by (simp add: sum.delta)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   604
  also from assms have "... \<noteq> 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   605
  finally show "(f * g) $ (subdegree f + subdegree g) \<noteq> 0" .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   606
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   607
  fix m assume m: "m < subdegree f + subdegree g"
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   608
  have "(f * g) $ m = (\<Sum>i=0..m. f$i * g$(m-i))" by (simp add: fps_mult_nth)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   609
  also have "... = (\<Sum>i=0..m. 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   610
  proof (rule sum.cong)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   611
    fix i assume "i \<in> {0..m}"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   612
    with m have "i < subdegree f \<or> m - i < subdegree g" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   613
    thus "f$i * g$(m-i) = 0" by (elim disjE) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   614
  qed auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   615
  finally show "(f * g) $ m = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   616
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   617
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   618
lemma subdegree_power [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   619
  "subdegree ((f :: ('a :: ring_1_no_zero_divisors) fps) ^ n) = n * subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   620
  by (cases "f = 0"; induction n) simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   621
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   622
lemma subdegree_uminus [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   623
  "subdegree (-(f::('a::group_add) fps)) = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   624
  by (simp add: subdegree_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   625
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   626
lemma subdegree_minus_commute [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   627
  "subdegree (f-(g::('a::group_add) fps)) = subdegree (g - f)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   628
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   629
  have "f - g = -(g - f)" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   630
  also have "subdegree ... = subdegree (g - f)" by (simp only: subdegree_uminus)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   631
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   632
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   633
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   634
lemma subdegree_add_ge:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   635
  assumes "f \<noteq> -(g :: ('a :: {group_add}) fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   636
  shows   "subdegree (f + g) \<ge> min (subdegree f) (subdegree g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   637
proof (rule subdegree_geI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   638
  from assms show "f + g \<noteq> 0" by (subst (asm) eq_neg_iff_add_eq_0)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   639
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   640
  fix i assume "i < min (subdegree f) (subdegree g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   641
  hence "f $ i = 0" and "g $ i = 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   642
  thus "(f + g) $ i = 0" by force
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   643
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   644
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   645
lemma subdegree_add_eq1:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   646
  assumes "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   647
  assumes "subdegree f < subdegree (g :: ('a :: {group_add}) fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   648
  shows   "subdegree (f + g) = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   649
proof (rule antisym[OF subdegree_leI])
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   650
  from assms show "subdegree (f + g) \<ge> subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   651
    by (intro order.trans[OF min.boundedI subdegree_add_ge]) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   652
  from assms have "f $ subdegree f \<noteq> 0" "g $ subdegree f = 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   653
  thus "(f + g) $ subdegree f \<noteq> 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   654
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   655
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   656
lemma subdegree_add_eq2:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   657
  assumes "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   658
  assumes "subdegree g < subdegree (f :: ('a :: {ab_group_add}) fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   659
  shows   "subdegree (f + g) = subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   660
  using subdegree_add_eq1[OF assms] by (simp add: add.commute)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   661
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   662
lemma subdegree_diff_eq1:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   663
  assumes "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   664
  assumes "subdegree f < subdegree (g :: ('a :: {ab_group_add}) fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   665
  shows   "subdegree (f - g) = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   666
  using subdegree_add_eq1[of f "-g"] assms by (simp add: add.commute)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   667
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   668
lemma subdegree_diff_eq2:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   669
  assumes "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   670
  assumes "subdegree g < subdegree (f :: ('a :: {ab_group_add}) fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   671
  shows   "subdegree (f - g) = subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   672
  using subdegree_add_eq2[of "-g" f] assms by (simp add: add.commute)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   673
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   674
lemma subdegree_diff_ge [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   675
  assumes "f \<noteq> (g :: ('a :: {group_add}) fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   676
  shows   "subdegree (f - g) \<ge> min (subdegree f) (subdegree g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   677
  using assms subdegree_add_ge[of f "-g"] by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   678
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   679
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   680
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   681
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   682
subsection \<open>Shifting and slicing\<close>
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   683
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   684
definition fps_shift :: "nat \<Rightarrow> 'a fps \<Rightarrow> 'a fps" where
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   685
  "fps_shift n f = Abs_fps (\<lambda>i. f $ (i + n))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   686
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   687
lemma fps_shift_nth [simp]: "fps_shift n f $ i = f $ (i + n)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   688
  by (simp add: fps_shift_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   689
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   690
lemma fps_shift_0 [simp]: "fps_shift 0 f = f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   691
  by (intro fps_ext) (simp add: fps_shift_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   692
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   693
lemma fps_shift_zero [simp]: "fps_shift n 0 = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   694
  by (intro fps_ext) (simp add: fps_shift_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   695
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   696
lemma fps_shift_one: "fps_shift n 1 = (if n = 0 then 1 else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   697
  by (intro fps_ext) (simp add: fps_shift_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   698
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   699
lemma fps_shift_fps_const: "fps_shift n (fps_const c) = (if n = 0 then fps_const c else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   700
  by (intro fps_ext) (simp add: fps_shift_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   701
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   702
lemma fps_shift_numeral: "fps_shift n (numeral c) = (if n = 0 then numeral c else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   703
  by (simp add: numeral_fps_const fps_shift_fps_const)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   704
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   705
lemma fps_shift_fps_X_power [simp]:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   706
  "n \<le> m \<Longrightarrow> fps_shift n (fps_X ^ m) = (fps_X ^ (m - n) ::'a::comm_ring_1 fps)"
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   707
  by (intro fps_ext) (auto simp: fps_shift_def )
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   708
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   709
lemma fps_shift_times_fps_X_power:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   710
  "n \<le> subdegree f \<Longrightarrow> fps_shift n f * fps_X ^ n = (f :: 'a :: comm_ring_1 fps)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   711
  by (intro fps_ext) (auto simp: fps_X_power_mult_right_nth nth_less_subdegree_zero)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   712
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   713
lemma fps_shift_times_fps_X_power' [simp]:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   714
  "fps_shift n (f * fps_X^n) = (f :: 'a :: comm_ring_1 fps)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   715
  by (intro fps_ext) (auto simp: fps_X_power_mult_right_nth nth_less_subdegree_zero)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   716
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   717
lemma fps_shift_times_fps_X_power'':
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   718
  "m \<le> n \<Longrightarrow> fps_shift n (f * fps_X^m) = fps_shift (n - m) (f :: 'a :: comm_ring_1 fps)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   719
  by (intro fps_ext) (auto simp: fps_X_power_mult_right_nth nth_less_subdegree_zero)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   720
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   721
lemma fps_shift_subdegree [simp]:
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   722
  "n \<le> subdegree f \<Longrightarrow> subdegree (fps_shift n f) = subdegree (f :: 'a :: comm_ring_1 fps) - n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   723
  by (cases "f = 0") (force intro: nth_less_subdegree_zero subdegreeI)+
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   724
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   725
lemma subdegree_decompose:
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   726
  "f = fps_shift (subdegree f) f * fps_X ^ subdegree (f :: ('a :: comm_ring_1) fps)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   727
  by (rule fps_ext) (auto simp: fps_X_power_mult_right_nth)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   728
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   729
lemma subdegree_decompose':
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   730
  "n \<le> subdegree (f :: ('a :: comm_ring_1) fps) \<Longrightarrow> f = fps_shift n f * fps_X^n"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   731
  by (rule fps_ext) (auto simp: fps_X_power_mult_right_nth intro!: nth_less_subdegree_zero)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   732
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   733
lemma fps_shift_fps_shift:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   734
  "fps_shift (m + n) f = fps_shift m (fps_shift n f)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   735
  by (rule fps_ext) (simp add: add_ac)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   736
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   737
lemma fps_shift_add:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   738
  "fps_shift n (f + g) = fps_shift n f + fps_shift n g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   739
  by (simp add: fps_eq_iff)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   740
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   741
lemma fps_shift_mult:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   742
  assumes "n \<le> subdegree (g :: 'b :: {comm_ring_1} fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   743
  shows   "fps_shift n (h*g) = h * fps_shift n g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   744
proof -
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   745
  from assms have "g = fps_shift n g * fps_X^n" by (rule subdegree_decompose')
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   746
  also have "h * ... = (h * fps_shift n g) * fps_X^n" by simp
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   747
  also have "fps_shift n ... = h * fps_shift n g" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   748
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   749
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   750
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   751
lemma fps_shift_mult_right:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   752
  assumes "n \<le> subdegree (g :: 'b :: {comm_ring_1} fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   753
  shows   "fps_shift n (g*h) = h * fps_shift n g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   754
  by (subst mult.commute, subst fps_shift_mult) (simp_all add: assms)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   755
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   756
lemma nth_subdegree_zero_iff [simp]: "f $ subdegree f = 0 \<longleftrightarrow> f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   757
  by (cases "f = 0") auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   758
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   759
lemma fps_shift_subdegree_zero_iff [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   760
  "fps_shift (subdegree f) f = 0 \<longleftrightarrow> f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   761
  by (subst (1) nth_subdegree_zero_iff[symmetric], cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   762
     (simp_all del: nth_subdegree_zero_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   763
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   764
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   765
definition "fps_cutoff n f = Abs_fps (\<lambda>i. if i < n then f$i else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   766
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   767
lemma fps_cutoff_nth [simp]: "fps_cutoff n f $ i = (if i < n then f$i else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   768
  unfolding fps_cutoff_def by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   769
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   770
lemma fps_cutoff_zero_iff: "fps_cutoff n f = 0 \<longleftrightarrow> (f = 0 \<or> n \<le> subdegree f)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   771
proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   772
  assume A: "fps_cutoff n f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   773
  thus "f = 0 \<or> n \<le> subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   774
  proof (cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   775
    assume "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   776
    with A have "n \<le> subdegree f"
62390
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
   777
      by (intro subdegree_geI) (auto simp: fps_eq_iff split: if_split_asm)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   778
    thus ?thesis ..
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   779
  qed simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   780
qed (auto simp: fps_eq_iff intro: nth_less_subdegree_zero)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   781
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   782
lemma fps_cutoff_0 [simp]: "fps_cutoff 0 f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   783
  by (simp add: fps_eq_iff)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   784
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   785
lemma fps_cutoff_zero [simp]: "fps_cutoff n 0 = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   786
  by (simp add: fps_eq_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   787
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   788
lemma fps_cutoff_one: "fps_cutoff n 1 = (if n = 0 then 0 else 1)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   789
  by (simp add: fps_eq_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   790
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   791
lemma fps_cutoff_fps_const: "fps_cutoff n (fps_const c) = (if n = 0 then 0 else fps_const c)"
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   792
  by (simp add: fps_eq_iff)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   793
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   794
lemma fps_cutoff_numeral: "fps_cutoff n (numeral c) = (if n = 0 then 0 else numeral c)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   795
  by (simp add: numeral_fps_const fps_cutoff_fps_const)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   796
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   797
lemma fps_shift_cutoff:
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   798
  "fps_shift n (f :: ('a :: comm_ring_1) fps) * fps_X^n + fps_cutoff n f = f"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   799
  by (simp add: fps_eq_iff fps_X_power_mult_right_nth)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   800
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   801
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   802
subsection \<open>Formal Power series form a metric space\<close>
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   803
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   804
instantiation fps :: (comm_ring_1) dist
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   805
begin
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   806
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   807
definition
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   808
  dist_fps_def: "dist (a :: 'a fps) b = (if a = b then 0 else inverse (2 ^ subdegree (a - b)))"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   809
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   810
lemma dist_fps_ge0: "dist (a :: 'a fps) b \<ge> 0"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   811
  by (simp add: dist_fps_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   812
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   813
lemma dist_fps_sym: "dist (a :: 'a fps) b = dist b a"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   814
  by (simp add: dist_fps_def)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   815
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   816
instance ..
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   817
30746
d6915b738bd9 fps made instance of number_ring
chaieb
parents: 30488
diff changeset
   818
end
d6915b738bd9 fps made instance of number_ring
chaieb
parents: 30488
diff changeset
   819
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   820
instantiation fps :: (comm_ring_1) metric_space
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   821
begin
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   822
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
   823
definition uniformity_fps_def [code del]:
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
   824
  "(uniformity :: ('a fps \<times> 'a fps) filter) = (INF e\<in>{0 <..}. principal {(x, y). dist x y < e})"
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
   825
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
   826
definition open_fps_def' [code del]:
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
   827
  "open (U :: 'a fps set) \<longleftrightarrow> (\<forall>x\<in>U. eventually (\<lambda>(x', y). x' = x \<longrightarrow> y \<in> U) uniformity)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   828
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   829
instance
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   830
proof
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   831
  show th: "dist a b = 0 \<longleftrightarrow> a = b" for a b :: "'a fps"
62390
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
   832
    by (simp add: dist_fps_def split: if_split_asm)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   833
  then have th'[simp]: "dist a a = 0" for a :: "'a fps" by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   834
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   835
  fix a b c :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   836
  consider "a = b" | "c = a \<or> c = b" | "a \<noteq> b" "a \<noteq> c" "b \<noteq> c" by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   837
  then show "dist a b \<le> dist a c + dist b c"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   838
  proof cases
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   839
    case 1
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   840
    then show ?thesis by (simp add: dist_fps_def)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   841
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   842
    case 2
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   843
    then show ?thesis
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   844
      by (cases "c = a") (simp_all add: th dist_fps_sym)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   845
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
   846
    case neq: 3
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   847
    have False if "dist a b > dist a c + dist b c"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   848
    proof -
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   849
      let ?n = "subdegree (a - b)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   850
      from neq have "dist a b > 0" "dist b c > 0" and "dist a c > 0" by (simp_all add: dist_fps_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   851
      with that have "dist a b > dist a c" and "dist a b > dist b c" by simp_all
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   852
      with neq have "?n < subdegree (a - c)" and "?n < subdegree (b - c)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   853
        by (simp_all add: dist_fps_def field_simps)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   854
      hence "(a - c) $ ?n = 0" and "(b - c) $ ?n = 0"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   855
        by (simp_all only: nth_less_subdegree_zero)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   856
      hence "(a - b) $ ?n = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   857
      moreover from neq have "(a - b) $ ?n \<noteq> 0" by (intro nth_subdegree_nonzero) simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   858
      ultimately show False by contradiction
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   859
    qed
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   860
    thus ?thesis by (auto simp add: not_le[symmetric])
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   861
  qed
62101
26c0a70f78a3 add uniform spaces
hoelzl
parents: 61969
diff changeset
   862
qed (rule open_fps_def' uniformity_fps_def)+
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   863
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   864
end
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   865
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   866
declare uniformity_Abort[where 'a="'a :: comm_ring_1 fps", code]
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   867
66373
56f8bfe1211c Removed unnecessary constant 'ball' from Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 66311
diff changeset
   868
lemma open_fps_def: "open (S :: 'a::comm_ring_1 fps set) = (\<forall>a \<in> S. \<exists>r. r >0 \<and> {y. dist y a < r} \<subseteq> S)"
56f8bfe1211c Removed unnecessary constant 'ball' from Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 66311
diff changeset
   869
  unfolding open_dist subset_eq by simp
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   870
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   871
text \<open>The infinite sums and justification of the notation in textbooks.\<close>
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   872
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   873
lemma reals_power_lt_ex:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   874
  fixes x y :: real
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   875
  assumes xp: "x > 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   876
    and y1: "y > 1"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   877
  shows "\<exists>k>0. (1/y)^k < x"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   878
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   879
  have yp: "y > 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   880
    using y1 by simp
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   881
  from reals_Archimedean2[of "max 0 (- log y x) + 1"]
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   882
  obtain k :: nat where k: "real k > max 0 (- log y x) + 1"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   883
    by blast
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   884
  from k have kp: "k > 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   885
    by simp
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   886
  from k have "real k > - log y x"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   887
    by simp
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   888
  then have "ln y * real k > - ln x"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   889
    unfolding log_def
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   890
    using ln_gt_zero_iff[OF yp] y1
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   891
    by (simp add: minus_divide_left field_simps del: minus_divide_left[symmetric])
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   892
  then have "ln y * real k + ln x > 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   893
    by simp
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   894
  then have "exp (real k * ln y + ln x) > exp 0"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
   895
    by (simp add: ac_simps)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   896
  then have "y ^ k * x > 1"
65578
e4997c181cce New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents: 65435
diff changeset
   897
    unfolding exp_zero exp_add exp_of_nat_mult exp_ln [OF xp] exp_ln [OF yp]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   898
    by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   899
  then have "x > (1 / y)^k" using yp
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   900
    by (simp add: field_simps)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   901
  then show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   902
    using kp by blast
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   903
qed
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   904
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   905
lemma fps_sum_rep_nth: "(sum (\<lambda>i. fps_const(a$i)*fps_X^i) {0..m})$n =
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   906
    (if n \<le> m then a$n else 0::'a::comm_ring_1)"
66089
def95e0bc529 Some new material. SIMPRULE STATUS for sum/prod.delta rules!
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
   907
  by (auto simp add: fps_sum_nth cond_value_iff cong del: if_weak_cong)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   908
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
   909
lemma fps_notation: "(\<lambda>n. sum (\<lambda>i. fps_const(a$i) * fps_X^i) {0..n}) \<longlonglongrightarrow> a"
61969
e01015e49041 more symbols;
wenzelm
parents: 61943
diff changeset
   910
  (is "?s \<longlonglongrightarrow> a")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   911
proof -
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   912
  have "\<exists>n0. \<forall>n \<ge> n0. dist (?s n) a < r" if "r > 0" for r
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   913
  proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   914
    obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   915
      using reals_power_lt_ex[OF \<open>r > 0\<close>, of 2] by auto
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   916
    show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   917
    proof -
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   918
      have "dist (?s n) a < r" if nn0: "n \<ge> n0" for n
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   919
      proof -
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   920
        from that have thnn0: "(1/2)^n \<le> (1/2 :: real)^n0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   921
          by (simp add: divide_simps)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   922
        show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   923
        proof (cases "?s n = a")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   924
          case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   925
          then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   926
            unfolding dist_eq_0_iff[of "?s n" a, symmetric]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   927
            using \<open>r > 0\<close> by (simp del: dist_eq_0_iff)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   928
        next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   929
          case False
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   930
          from False have dth: "dist (?s n) a = (1/2)^subdegree (?s n - a)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   931
            by (simp add: dist_fps_def field_simps)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   932
          from False have kn: "subdegree (?s n - a) > n"
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   933
            by (intro subdegree_greaterI) (simp_all add: fps_sum_rep_nth)
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   934
          then have "dist (?s n) a < (1/2)^n"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   935
            by (simp add: field_simps dist_fps_def)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   936
          also have "\<dots> \<le> (1/2)^n0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   937
            using nn0 by (simp add: divide_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   938
          also have "\<dots> < r"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   939
            using n0 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   940
          finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   941
        qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   942
      qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   943
      then show ?thesis by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   944
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   945
  qed
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   946
  then show ?thesis
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   947
    unfolding lim_sequentially by blast
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   948
qed
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   949
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   950
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   951
subsection \<open>Inverses of formal power series\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   952
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   953
declare sum.cong[fundef_cong]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   954
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   955
instantiation fps :: ("{comm_monoid_add,inverse,times,uminus}") inverse
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   956
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   957
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   958
fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   959
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   960
  "natfun_inverse f 0 = inverse (f$0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
   961
| "natfun_inverse f n = - inverse (f$0) * sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   962
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   963
definition fps_inverse_def: "inverse f = (if f $ 0 = 0 then 0 else Abs_fps (natfun_inverse f))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   964
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   965
definition fps_divide_def:
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
   966
  "f div g = (if g = 0 then 0 else
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   967
     let n = subdegree g; h = fps_shift n g
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   968
     in  fps_shift n (f * inverse h))"
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   969
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   970
instance ..
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   971
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   972
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   973
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   974
lemma fps_inverse_zero [simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   975
  "inverse (0 :: 'a::{comm_monoid_add,inverse,times,uminus} fps) = 0"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   976
  by (simp add: fps_ext fps_inverse_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   977
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   978
lemma fps_inverse_one [simp]: "inverse (1 :: 'a::{division_ring,zero_neq_one} fps) = 1"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   979
  apply (auto simp add: expand_fps_eq fps_inverse_def)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   980
  apply (case_tac n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   981
  apply auto
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   982
  done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   983
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   984
lemma inverse_mult_eq_1 [intro]:
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   985
  assumes f0: "f$0 \<noteq> (0::'a::field)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   986
  shows "inverse f * f = 1"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   987
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   988
  have c: "inverse f * f = f * inverse f"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   989
    by (simp add: mult.commute)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   990
  from f0 have ifn: "\<And>n. inverse f $ n = natfun_inverse f n"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   991
    by (simp add: fps_inverse_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   992
  from f0 have th0: "(inverse f * f) $ 0 = 1"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   993
    by (simp add: fps_mult_nth fps_inverse_def)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   994
  have "(inverse f * f)$n = 0" if np: "n > 0" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   995
  proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   996
    from np have eq: "{0..n} = {0} \<union> {1 .. n}"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   997
      by auto
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   998
    have d: "{0} \<inter> {1 .. n} = {}"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   999
      by auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1000
    from f0 np have th0: "- (inverse f $ n) =
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1001
      (sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1002
      by (cases n) (simp_all add: divide_inverse fps_inverse_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1003
    from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1004
    have th1: "sum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n} = - (f$0) * (inverse f)$n"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1005
      by (simp add: field_simps)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1006
    have "(f * inverse f) $ n = (\<Sum>i = 0..n. f $i * natfun_inverse f (n - i))"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1007
      unfolding fps_mult_nth ifn ..
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1008
    also have "\<dots> = f$0 * natfun_inverse f n + (\<Sum>i = 1..n. f$i * natfun_inverse f (n-i))"
46757
ad878aff9c15 removing finiteness goals
bulwahn
parents: 46131
diff changeset
  1009
      by (simp add: eq)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1010
    also have "\<dots> = 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1011
      unfolding th1 ifn by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1012
    finally show ?thesis unfolding c .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1013
  qed
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1014
  with th0 show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1015
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1016
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1017
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1018
lemma fps_inverse_0_iff[simp]: "(inverse f) $ 0 = (0::'a::division_ring) \<longleftrightarrow> f $ 0 = 0"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  1019
  by (simp add: fps_inverse_def nonzero_imp_inverse_nonzero)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1020
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1021
lemma fps_inverse_nth_0 [simp]: "inverse f $ 0 = inverse (f $ 0 :: 'a :: division_ring)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1022
  by (simp add: fps_inverse_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1023
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1024
lemma fps_inverse_eq_0_iff[simp]: "inverse f = (0:: ('a::division_ring) fps) \<longleftrightarrow> f $ 0 = 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1025
proof
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1026
  assume A: "inverse f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1027
  have "0 = inverse f $ 0" by (subst A) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1028
  thus "f $ 0 = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1029
qed (simp add: fps_inverse_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1030
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1031
lemma fps_inverse_idempotent[intro, simp]:
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1032
  assumes f0: "f$0 \<noteq> (0::'a::field)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1033
  shows "inverse (inverse f) = f"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1034
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1035
  from f0 have if0: "inverse f $ 0 \<noteq> 0" by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1036
  from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1037
  have "inverse f * f = inverse f * inverse (inverse f)"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1038
    by (simp add: ac_simps)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1039
  then show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1040
    using f0 unfolding mult_cancel_left by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1041
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1042
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1043
lemma fps_inverse_unique:
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1044
  assumes fg: "(f :: 'a :: field fps) * g = 1"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1045
  shows   "inverse f = g"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1046
proof -
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1047
  have f0: "f $ 0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1048
  proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1049
    assume "f $ 0 = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1050
    hence "0 = (f * g) $ 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1051
    also from fg have "(f * g) $ 0 = 1" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1052
    finally show False by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1053
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1054
  from inverse_mult_eq_1[OF this] fg
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1055
  have th0: "inverse f * f = g * f"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  1056
    by (simp add: ac_simps)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1057
  then show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1058
    using f0
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1059
    unfolding mult_cancel_right
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  1060
    by (auto simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1061
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1062
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1063
lemma fps_inverse_eq_0: "f$0 = 0 \<Longrightarrow> inverse (f :: 'a :: division_ring fps) = 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1064
  by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1065
  
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1066
lemma sum_zero_lemma:
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1067
  fixes n::nat
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1068
  assumes "0 < n"
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1069
  shows "(\<Sum>i = 0..n. if n = i then 1 else if n - i = 1 then - 1 else 0) = (0::'a::field)"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1070
proof -
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1071
  let ?f = "\<lambda>i. if n = i then 1 else if n - i = 1 then - 1 else 0"
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1072
  let ?g = "\<lambda>i. if i = n then 1 else if i = n - 1 then - 1 else 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1073
  let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1074
  have th1: "sum ?f {0..n} = sum ?g {0..n}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1075
    by (rule sum.cong) auto
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1076
  have th2: "sum ?g {0..n - 1} = sum ?h {0..n - 1}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1077
    apply (rule sum.cong)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1078
    using assms
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1079
    apply auto
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1080
    done
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1081
  have eq: "{0 .. n} = {0.. n - 1} \<union> {n}"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1082
    by auto
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1083
  from assms have d: "{0.. n - 1} \<inter> {n} = {}"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1084
    by auto
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1085
  have f: "finite {0.. n - 1}" "finite {n}"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1086
    by auto
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1087
  show ?thesis
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1088
    unfolding th1
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1089
    apply (simp add: sum.union_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1090
    unfolding th2
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1091
    apply (simp add: sum.delta)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1092
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1093
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1094
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1095
lemma fps_inverse_mult: "inverse (f * g :: 'a::field fps) = inverse f * inverse g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1096
proof (cases "f$0 = 0 \<or> g$0 = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1097
  assume "\<not>(f$0 = 0 \<or> g$0 = 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1098
  hence [simp]: "f$0 \<noteq> 0" "g$0 \<noteq> 0" by simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1099
  show ?thesis
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1100
  proof (rule fps_inverse_unique)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1101
    have "f * g * (inverse f * inverse g) = (inverse f * f) * (inverse g * g)" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1102
    also have "... = 1" by (subst (1 2) inverse_mult_eq_1) simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1103
    finally show "f * g * (inverse f * inverse g) = 1" .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1104
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1105
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1106
  assume A: "f$0 = 0 \<or> g$0 = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1107
  hence "inverse (f * g) = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1108
  also from A have "... = inverse f * inverse g" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1109
  finally show "inverse (f * g) = inverse f * inverse g" .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1110
qed
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1111
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1112
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1113
lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field))) =
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1114
    Abs_fps (\<lambda>n. if n= 0 then 1 else if n=1 then - 1 else 0)"
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1115
  apply (rule fps_inverse_unique)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1116
  apply (simp_all add: fps_eq_iff fps_mult_nth sum_zero_lemma)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1117
  done
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1118
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1119
lemma subdegree_inverse [simp]: "subdegree (inverse (f::'a::field fps)) = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1120
proof (cases "f$0 = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1121
  assume nz: "f$0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1122
  hence "subdegree (inverse f) + subdegree f = subdegree (inverse f * f)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1123
    by (subst subdegree_mult) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1124
  also from nz have "subdegree f = 0" by (simp add: subdegree_eq_0_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1125
  also from nz have "inverse f * f = 1" by (rule inverse_mult_eq_1)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1126
  finally show "subdegree (inverse f) = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1127
qed (simp_all add: fps_inverse_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1128
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1129
lemma fps_is_unit_iff [simp]: "(f :: 'a :: field fps) dvd 1 \<longleftrightarrow> f $ 0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1130
proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1131
  assume "f dvd 1"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1132
  then obtain g where "1 = f * g" by (elim dvdE)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1133
  from this[symmetric] have "(f*g) $ 0 = 1" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1134
  thus "f $ 0 \<noteq> 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1135
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1136
  assume A: "f $ 0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1137
  thus "f dvd 1" by (simp add: inverse_mult_eq_1[OF A, symmetric])
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1138
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1139
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1140
lemma subdegree_eq_0' [simp]: "(f :: 'a :: field fps) dvd 1 \<Longrightarrow> subdegree f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1141
  by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1142
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1143
lemma fps_unit_dvd [simp]: "(f $ 0 :: 'a :: field) \<noteq> 0 \<Longrightarrow> f dvd g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1144
  by (rule dvd_trans, subst fps_is_unit_iff) simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1145
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1146
instantiation fps :: (field) normalization_semidom
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1147
begin
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1148
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1149
definition fps_unit_factor_def [simp]:
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1150
  "unit_factor f = fps_shift (subdegree f) f"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1151
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1152
definition fps_normalize_def [simp]:
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1153
  "normalize f = (if f = 0 then 0 else fps_X ^ subdegree f)"
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1154
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1155
instance proof
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1156
  fix f :: "'a fps"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1157
  show "unit_factor f * normalize f = f"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1158
    by (simp add: fps_shift_times_fps_X_power)
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1159
next
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1160
  fix f g :: "'a fps"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1161
  show "unit_factor (f * g) = unit_factor f * unit_factor g"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1162
  proof (cases "f = 0 \<or> g = 0")
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1163
    assume "\<not>(f = 0 \<or> g = 0)"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1164
    thus "unit_factor (f * g) = unit_factor f * unit_factor g"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1165
    unfolding fps_unit_factor_def
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1166
      by (auto simp: fps_shift_fps_shift fps_shift_mult fps_shift_mult_right)
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1167
  qed auto
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1168
next
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1169
  fix f g :: "'a fps"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1170
  assume "g \<noteq> 0"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1171
  then have "f * (fps_shift (subdegree g) g * inverse (fps_shift (subdegree g) g)) = f"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1172
    by (metis add_cancel_right_left fps_shift_nth inverse_mult_eq_1 mult.commute mult_cancel_left2 nth_subdegree_nonzero)
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1173
  then have "fps_shift (subdegree g) (g * (f * inverse (fps_shift (subdegree g) g))) = f"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1174
    by (simp add: fps_shift_mult_right mult.commute)
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1175
  with \<open>g \<noteq> 0\<close> show "f * g / g = f"
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1176
    by (simp add: fps_divide_def Let_def ac_simps)
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1177
qed (auto simp add: fps_divide_def Let_def)
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1178
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1179
end
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1180
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1181
instantiation fps :: (field) idom_modulo
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1182
begin
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1183
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1184
definition fps_mod_def:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1185
  "f mod g = (if g = 0 then f else
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1186
     let n = subdegree g; h = fps_shift n g
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1187
     in  fps_cutoff n (f * inverse h) * h)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1188
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1189
lemma fps_mod_eq_zero:
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1190
  assumes "g \<noteq> 0" and "subdegree f \<ge> subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1191
  shows   "f mod g = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1192
  using assms by (cases "f = 0") (auto simp: fps_cutoff_zero_iff fps_mod_def Let_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1193
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1194
lemma fps_times_divide_eq:
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1195
  assumes "g \<noteq> 0" and "subdegree f \<ge> subdegree (g :: 'a fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1196
  shows   "f div g * g = f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1197
proof (cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1198
  assume nz: "f \<noteq> 0"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62481
diff changeset
  1199
  define n where "n = subdegree g"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62481
diff changeset
  1200
  define h where "h = fps_shift n g"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1201
  from assms have [simp]: "h $ 0 \<noteq> 0" unfolding h_def by (simp add: n_def)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1202
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1203
  from assms nz have "f div g * g = fps_shift n (f * inverse h) * g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1204
    by (simp add: fps_divide_def Let_def h_def n_def)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1205
  also have "... = fps_shift n (f * inverse h) * fps_X^n * h" unfolding h_def n_def
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1206
    by (subst subdegree_decompose[of g]) simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1207
  also have "fps_shift n (f * inverse h) * fps_X^n = f * inverse h"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1208
    by (rule fps_shift_times_fps_X_power) (simp_all add: nz assms n_def)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1209
  also have "... * h = f * (inverse h * h)" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1210
  also have "inverse h * h = 1" by (rule inverse_mult_eq_1) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1211
  finally show ?thesis by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1212
qed (simp_all add: fps_divide_def Let_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1213
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1214
lemma
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1215
  assumes "g$0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1216
  shows   fps_divide_unit: "f div g = f * inverse g" and fps_mod_unit [simp]: "f mod g = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1217
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1218
  from assms have [simp]: "subdegree g = 0" by (simp add: subdegree_eq_0_iff)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1219
  from assms show "f div g = f * inverse g"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1220
    by (auto simp: fps_divide_def Let_def subdegree_eq_0_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1221
  from assms show "f mod g = 0" by (intro fps_mod_eq_zero) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1222
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1223
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1224
instance proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1225
  fix f g :: "'a fps"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62481
diff changeset
  1226
  define n where "n = subdegree g"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62481
diff changeset
  1227
  define h where "h = fps_shift n g"
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1228
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1229
  show "f div g * g + f mod g = f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1230
  proof (cases "g = 0 \<or> f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1231
    assume "\<not>(g = 0 \<or> f = 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1232
    hence nz [simp]: "f \<noteq> 0" "g \<noteq> 0" by simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1233
    show ?thesis
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1234
    proof (rule disjE[OF le_less_linear])
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1235
      assume "subdegree f \<ge> subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1236
      with nz show ?thesis by (simp add: fps_mod_eq_zero fps_times_divide_eq)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1237
    next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1238
      assume "subdegree f < subdegree g"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1239
      have g_decomp: "g = h * fps_X^n" unfolding h_def n_def by (rule subdegree_decompose)
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1240
      have "f div g * g + f mod g =
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1241
              fps_shift n (f * inverse h) * g + fps_cutoff n (f * inverse h) * h"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1242
        by (simp add: fps_mod_def fps_divide_def Let_def n_def h_def)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1243
      also have "... = h * (fps_shift n (f * inverse h) * fps_X^n + fps_cutoff n (f * inverse h))"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1244
        by (subst g_decomp) (simp add: algebra_simps)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1245
      also have "... = f * (inverse h * h)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1246
        by (subst fps_shift_cutoff) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1247
      also have "inverse h * h = 1" by (rule inverse_mult_eq_1) (simp add: h_def n_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1248
      finally show ?thesis by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1249
    qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1250
  qed (auto simp: fps_mod_def fps_divide_def Let_def)
64592
7759f1766189 more fine-grained type class hierarchy for div and mod
haftmann
parents: 64272
diff changeset
  1251
qed
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1252
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1253
end
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1254
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1255
lemma subdegree_mod:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1256
  assumes "f \<noteq> 0" "subdegree f < subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1257
  shows   "subdegree (f mod g) = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1258
proof (cases "f div g * g = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1259
  assume "f div g * g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1260
  hence [simp]: "f div g \<noteq> 0" "g \<noteq> 0" by auto
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1261
  from div_mult_mod_eq[of f g] have "f mod g = f - f div g * g" by (simp add: algebra_simps)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1262
  also from assms have "subdegree ... = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1263
    by (intro subdegree_diff_eq1) simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1264
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1265
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1266
  assume zero: "f div g * g = 0"
64242
93c6f0da5c70 more standardized theorem names for facts involving the div and mod identity
haftmann
parents: 64240
diff changeset
  1267
  from div_mult_mod_eq[of f g] have "f mod g = f - f div g * g" by (simp add: algebra_simps)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1268
  also note zero
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1269
  finally show ?thesis by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1270
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1271
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1272
lemma fps_divide_nth_0 [simp]: "g $ 0 \<noteq> 0 \<Longrightarrow> (f div g) $ 0 = f $ 0 / (g $ 0 :: _ :: field)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1273
  by (simp add: fps_divide_unit divide_inverse)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1274
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1275
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1276
lemma dvd_imp_subdegree_le:
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1277
  "(f :: 'a :: idom fps) dvd g \<Longrightarrow> g \<noteq> 0 \<Longrightarrow> subdegree f \<le> subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1278
  by (auto elim: dvdE)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1279
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1280
lemma fps_dvd_iff:
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1281
  assumes "(f :: 'a :: field fps) \<noteq> 0" "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1282
  shows   "f dvd g \<longleftrightarrow> subdegree f \<le> subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1283
proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1284
  assume "subdegree f \<le> subdegree g"
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1285
  with assms have "g mod f = 0"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1286
    by (simp add: fps_mod_def Let_def fps_cutoff_zero_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1287
  thus "f dvd g" by (simp add: dvd_eq_mod_eq_0)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1288
qed (simp add: assms dvd_imp_subdegree_le)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1289
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1290
lemma fps_shift_altdef:
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1291
  "fps_shift n f = (f :: 'a :: field fps) div fps_X^n"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1292
  by (simp add: fps_divide_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1293
  
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1294
lemma fps_div_fps_X_power_nth: "((f :: 'a :: field fps) div fps_X^n) $ k = f $ (k + n)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1295
  by (simp add: fps_shift_altdef [symmetric])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1296
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1297
lemma fps_div_fps_X_nth: "((f :: 'a :: field fps) div fps_X) $ k = f $ Suc k"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1298
  using fps_div_fps_X_power_nth[of f 1] by simp
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1299
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1300
lemma fps_const_inverse: "inverse (fps_const (a::'a::field)) = fps_const (inverse a)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1301
  by (cases "a \<noteq> 0", rule fps_inverse_unique) (auto simp: fps_eq_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1302
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1303
lemma fps_const_divide: "fps_const (x :: _ :: field) / fps_const y = fps_const (x / y)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1304
  by (cases "y = 0") (simp_all add: fps_divide_unit fps_const_inverse divide_inverse)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1305
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1306
lemma inverse_fps_numeral:
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1307
  "inverse (numeral n :: ('a :: field_char_0) fps) = fps_const (inverse (numeral n))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1308
  by (intro fps_inverse_unique fps_ext) (simp_all add: fps_numeral_nth)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1309
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1310
lemma fps_numeral_divide_divide:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1311
  "x / numeral b / numeral c = (x / numeral (b * c) :: 'a :: field fps)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1312
  by (cases "numeral b = (0::'a)"; cases "numeral c = (0::'a)")
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1313
      (simp_all add: fps_divide_unit fps_inverse_mult [symmetric] numeral_fps_const numeral_mult 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1314
                del: numeral_mult [symmetric])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1315
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1316
lemma fps_numeral_mult_divide:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1317
  "numeral b * x / numeral c = (numeral b / numeral c * x :: 'a :: field fps)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1318
  by (cases "numeral c = (0::'a)") (simp_all add: fps_divide_unit numeral_fps_const)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1319
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1320
lemmas fps_numeral_simps = 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1321
  fps_numeral_divide_divide fps_numeral_mult_divide inverse_fps_numeral neg_numeral_fps_const
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1322
66550
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1323
lemma subdegree_div:
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1324
  assumes "q dvd p"
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1325
  shows   "subdegree ((p :: 'a :: field fps) div q) = subdegree p - subdegree q"
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1326
proof (cases "p = 0")
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1327
  case False
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1328
  from assms have "p = p div q * q" by simp
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1329
  also from assms False have "subdegree \<dots> = subdegree (p div q) + subdegree q"
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1330
    by (intro subdegree_mult) (auto simp: dvd_div_eq_0_iff)
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1331
  finally show ?thesis by simp
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1332
qed simp_all
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1333
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1334
lemma subdegree_div_unit:
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1335
  assumes "q $ 0 \<noteq> 0"
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1336
  shows   "subdegree ((p :: 'a :: field fps) div q) = subdegree p"
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1337
  using assms by (subst subdegree_div) simp_all
e5d82cf3c387 Some small lemmas about polynomials and FPSs
eberlm <eberlm@in.tum.de>
parents: 66480
diff changeset
  1338
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1339
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1340
subsection \<open>Formal power series form a Euclidean ring\<close>
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1341
64784
5cb5e7ecb284 reshaped euclidean semiring into hierarchy of euclidean semirings culminating in uniquely determined euclidean divion
haftmann
parents: 64592
diff changeset
  1342
instantiation fps :: (field) euclidean_ring_cancel
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1343
begin
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1344
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1345
definition fps_euclidean_size_def:
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62390
diff changeset
  1346
  "euclidean_size f = (if f = 0 then 0 else 2 ^ subdegree f)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1347
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1348
context
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1349
begin
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1350
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1351
private lemma fps_divide_cancel_aux1:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1352
  assumes "h$0 \<noteq> (0 :: 'a :: field)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1353
  shows   "(h * f) div (h * g) = f div g"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1354
proof (cases "g = 0")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1355
  assume "g \<noteq> 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1356
  from assms have "h \<noteq> 0" by auto
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1357
  note nz [simp] = \<open>g \<noteq> 0\<close> \<open>h \<noteq> 0\<close>
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1358
  from assms have [simp]: "subdegree h = 0" by (simp add: subdegree_eq_0_iff)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1359
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1360
  have "(h * f) div (h * g) =
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1361
          fps_shift (subdegree g) (h * f * inverse (fps_shift (subdegree g) (h*g)))"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1362
    by (simp add: fps_divide_def Let_def)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1363
  also have "h * f * inverse (fps_shift (subdegree g) (h*g)) =
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1364
               (inverse h * h) * f * inverse (fps_shift (subdegree g) g)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1365
    by (subst fps_shift_mult) (simp_all add: algebra_simps fps_inverse_mult)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1366
  also from assms have "inverse h * h = 1" by (rule inverse_mult_eq_1)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1367
  finally show "(h * f) div (h * g) = f div g" by (simp_all add: fps_divide_def Let_def)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1368
qed (simp_all add: fps_divide_def)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1369
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1370
private lemma fps_divide_cancel_aux2:
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1371
  "(f * fps_X^m) div (g * fps_X^m) = f div (g :: 'a :: field fps)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1372
proof (cases "g = 0")
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1373
  assume [simp]: "g \<noteq> 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1374
  have "(f * fps_X^m) div (g * fps_X^m) =
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1375
          fps_shift (subdegree g + m) (f*inverse (fps_shift (subdegree g + m) (g*fps_X^m))*fps_X^m)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1376
    by (simp add: fps_divide_def Let_def algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1377
  also have "... = f div g"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1378
    by (simp add: fps_shift_times_fps_X_power'' fps_divide_def Let_def)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1379
  finally show ?thesis .
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1380
qed (simp_all add: fps_divide_def)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1381
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1382
instance proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1383
  fix f g :: "'a fps" assume [simp]: "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1384
  show "euclidean_size f \<le> euclidean_size (f * g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1385
    by (cases "f = 0") (auto simp: fps_euclidean_size_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1386
  show "euclidean_size (f mod g) < euclidean_size g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1387
    apply (cases "f = 0", simp add: fps_euclidean_size_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1388
    apply (rule disjE[OF le_less_linear[of "subdegree g" "subdegree f"]])
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1389
    apply (simp_all add: fps_mod_eq_zero fps_euclidean_size_def subdegree_mod)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1390
    done
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1391
  show "(h * f) div (h * g) = f div g" if "h \<noteq> 0"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1392
    for f g h :: "'a fps"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1393
  proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1394
    define m where "m = subdegree h"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1395
    define h' where "h' = fps_shift m h"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1396
    have h_decomp: "h = h' * fps_X ^ m" unfolding h'_def m_def by (rule subdegree_decompose)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1397
    from \<open>h \<noteq> 0\<close> have [simp]: "h'$0 \<noteq> 0" by (simp add: h'_def m_def)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1398
    have "(h * f) div (h * g) = (h' * f * fps_X^m) div (h' * g * fps_X^m)"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1399
      by (simp add: h_decomp algebra_simps)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1400
    also have "... = f div g"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1401
      by (simp add: fps_divide_cancel_aux1 fps_divide_cancel_aux2)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1402
    finally show ?thesis .
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1403
  qed
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1404
  show "(f + g * h) div h = g + f div h"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1405
    if "h \<noteq> 0" for f g h :: "'a fps"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1406
  proof -
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1407
    define n h' where dfs: "n = subdegree h" "h' = fps_shift n h"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1408
    have "(f + g * h) div h = fps_shift n (f * inverse h') + fps_shift n (g * (h * inverse h'))"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1409
      by (simp add: fps_divide_def Let_def dfs [symmetric] algebra_simps fps_shift_add that)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1410
    also have "h * inverse h' = (inverse h' * h') * fps_X^n"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1411
      by (subst subdegree_decompose) (simp_all add: dfs)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1412
    also have "... = fps_X^n"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1413
      by (subst inverse_mult_eq_1) (simp_all add: dfs that)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1414
    also have "fps_shift n (g * fps_X^n) = g" by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1415
    also have "fps_shift n (f * inverse h') = f div h"
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1416
      by (simp add: fps_divide_def Let_def dfs)
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1417
    finally show ?thesis by simp
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1418
  qed
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62390
diff changeset
  1419
qed (simp_all add: fps_euclidean_size_def)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1420
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1421
end
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1422
66806
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1423
end
a4e82b58d833 abolished (semi)ring_div in favour of euclidean_(semi)ring_cancel
haftmann
parents: 66804
diff changeset
  1424
66817
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 66806
diff changeset
  1425
instance fps :: (field) normalization_euclidean_semiring ..
0b12755ccbb2 euclidean rings need no normalization
haftmann
parents: 66806
diff changeset
  1426
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1427
instantiation fps :: (field) euclidean_ring_gcd
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1428
begin
64786
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64784
diff changeset
  1429
definition fps_gcd_def: "(gcd :: 'a fps \<Rightarrow> _) = Euclidean_Algorithm.gcd"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64784
diff changeset
  1430
definition fps_lcm_def: "(lcm :: 'a fps \<Rightarrow> _) = Euclidean_Algorithm.lcm"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64784
diff changeset
  1431
definition fps_Gcd_def: "(Gcd :: 'a fps set \<Rightarrow> _) = Euclidean_Algorithm.Gcd"
340db65fd2c1 reworked to provide auxiliary operations Euclidean_Algorithm.* to instantiate gcd etc. for euclidean rings
haftmann
parents: 64784
diff changeset
  1432
definition fps_Lcm_def: "(Lcm :: 'a fps set \<Rightarrow> _) = Euclidean_Algorithm.Lcm"
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62390
diff changeset
  1433
instance by standard (simp_all add: fps_gcd_def fps_lcm_def fps_Gcd_def fps_Lcm_def)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1434
end
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1435
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1436
lemma fps_gcd:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1437
  assumes [simp]: "f \<noteq> 0" "g \<noteq> 0"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1438
  shows   "gcd f g = fps_X ^ min (subdegree f) (subdegree g)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1439
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1440
  let ?m = "min (subdegree f) (subdegree g)"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1441
  show "gcd f g = fps_X ^ ?m"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1442
  proof (rule sym, rule gcdI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1443
    fix d assume "d dvd f" "d dvd g"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1444
    thus "d dvd fps_X ^ ?m" by (cases "d = 0") (auto simp: fps_dvd_iff)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1445
  qed (simp_all add: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1446
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1447
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1448
lemma fps_gcd_altdef: "gcd (f :: 'a :: field fps) g =
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1449
  (if f = 0 \<and> g = 0 then 0 else
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1450
   if f = 0 then fps_X ^ subdegree g else
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1451
   if g = 0 then fps_X ^ subdegree f else
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1452
     fps_X ^ min (subdegree f) (subdegree g))"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1453
  by (simp add: fps_gcd)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1454
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1455
lemma fps_lcm:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1456
  assumes [simp]: "f \<noteq> 0" "g \<noteq> 0"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1457
  shows   "lcm f g = fps_X ^ max (subdegree f) (subdegree g)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1458
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1459
  let ?m = "max (subdegree f) (subdegree g)"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1460
  show "lcm f g = fps_X ^ ?m"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1461
  proof (rule sym, rule lcmI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1462
    fix d assume "f dvd d" "g dvd d"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1463
    thus "fps_X ^ ?m dvd d" by (cases "d = 0") (auto simp: fps_dvd_iff)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1464
  qed (simp_all add: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1465
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1466
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1467
lemma fps_lcm_altdef: "lcm (f :: 'a :: field fps) g =
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1468
  (if f = 0 \<or> g = 0 then 0 else fps_X ^ max (subdegree f) (subdegree g))"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1469
  by (simp add: fps_lcm)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1470
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1471
lemma fps_Gcd:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1472
  assumes "A - {0} \<noteq> {}"
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1473
  shows   "Gcd A = fps_X ^ (INF f\<in>A-{0}. subdegree f)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1474
proof (rule sym, rule GcdI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1475
  fix f assume "f \<in> A"
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1476
  thus "fps_X ^ (INF f\<in>A - {0}. subdegree f) dvd f"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1477
    by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cINF_lower)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1478
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1479
  fix d assume d: "\<And>f. f \<in> A \<Longrightarrow> d dvd f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1480
  from assms obtain f where "f \<in> A - {0}" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1481
  with d[of f] have [simp]: "d \<noteq> 0" by auto
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1482
  from d assms have "subdegree d \<le> (INF f\<in>A-{0}. subdegree f)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1483
    by (intro cINF_greatest) (auto simp: fps_dvd_iff[symmetric])
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1484
  with d assms show "d dvd fps_X ^ (INF f\<in>A-{0}. subdegree f)" by (simp add: fps_dvd_iff)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1485
qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1486
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1487
lemma fps_Gcd_altdef: "Gcd (A :: 'a :: field fps set) =
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1488
  (if A \<subseteq> {0} then 0 else fps_X ^ (INF f\<in>A-{0}. subdegree f))"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1489
  using fps_Gcd by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1490
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1491
lemma fps_Lcm:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1492
  assumes "A \<noteq> {}" "0 \<notin> A" "bdd_above (subdegree`A)"
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1493
  shows   "Lcm A = fps_X ^ (SUP f\<in>A. subdegree f)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1494
proof (rule sym, rule LcmI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1495
  fix f assume "f \<in> A"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1496
  moreover from assms(3) have "bdd_above (subdegree ` A)" by auto
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1497
  ultimately show "f dvd fps_X ^ (SUP f\<in>A. subdegree f)" using assms(2)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1498
    by (cases "f = 0") (auto simp: fps_dvd_iff intro!: cSUP_upper)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1499
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1500
  fix d assume d: "\<And>f. f \<in> A \<Longrightarrow> f dvd d"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1501
  from assms obtain f where f: "f \<in> A" "f \<noteq> 0" by auto
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1502
  show "fps_X ^ (SUP f\<in>A. subdegree f) dvd d"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1503
  proof (cases "d = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1504
    assume "d \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1505
    moreover from d have "\<And>f. f \<in> A \<Longrightarrow> f \<noteq> 0 \<Longrightarrow> f dvd d" by blast
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1506
    ultimately have "subdegree d \<ge> (SUP f\<in>A. subdegree f)" using assms
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1507
      by (intro cSUP_least) (auto simp: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1508
    with \<open>d \<noteq> 0\<close> show ?thesis by (simp add: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1509
  qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1510
qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1511
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1512
lemma fps_Lcm_altdef:
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  1513
  "Lcm (A :: 'a :: field fps set) =
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1514
     (if 0 \<in> A \<or> \<not>bdd_above (subdegree`A) then 0 else
69260
0a9688695a1b removed relics of ASCII syntax for indexed big operators
haftmann
parents: 69085
diff changeset
  1515
      if A = {} then 1 else fps_X ^ (SUP f\<in>A. subdegree f))"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1516
proof (cases "bdd_above (subdegree`A)")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1517
  assume unbounded: "\<not>bdd_above (subdegree`A)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1518
  have "Lcm A = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1519
  proof (rule ccontr)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1520
    assume "Lcm A \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1521
    from unbounded obtain f where f: "f \<in> A" "subdegree (Lcm A) < subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1522
      unfolding bdd_above_def by (auto simp: not_le)
63539
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63417
diff changeset
  1523
    moreover from f and \<open>Lcm A \<noteq> 0\<close> have "subdegree f \<le> subdegree (Lcm A)"
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62390
diff changeset
  1524
      by (intro dvd_imp_subdegree_le dvd_Lcm) simp_all
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1525
    ultimately show False by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1526
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1527
  with unbounded show ?thesis by simp
62422
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62390
diff changeset
  1528
qed (simp_all add: fps_Lcm Lcm_eq_0_I)
4aa35fd6c152 Tuned Euclidean rings
eberlm
parents: 62390
diff changeset
  1529
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1530
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1531
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1532
subsection \<open>Formal Derivatives, and the MacLaurin theorem around 0\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1533
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1534
definition "fps_deriv f = Abs_fps (\<lambda>n. of_nat (n + 1) * f $ (n + 1))"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1535
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1536
lemma fps_deriv_nth[simp]: "fps_deriv f $ n = of_nat (n +1) * f $ (n + 1)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1537
  by (simp add: fps_deriv_def)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1538
65398
eberlm <eberlm@in.tum.de>
parents: 65396
diff changeset
  1539
lemma fps_0th_higher_deriv: 
eberlm <eberlm@in.tum.de>
parents: 65396
diff changeset
  1540
  "(fps_deriv ^^ n) f $ 0 = (fact n * f $ n :: 'a :: {comm_ring_1, semiring_char_0})"
eberlm <eberlm@in.tum.de>
parents: 65396
diff changeset
  1541
  by (induction n arbitrary: f) (simp_all del: funpow.simps add: funpow_Suc_right algebra_simps)
eberlm <eberlm@in.tum.de>
parents: 65396
diff changeset
  1542
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1543
lemma fps_deriv_linear[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1544
  "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1545
    fps_const a * fps_deriv f + fps_const b * fps_deriv g"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1546
  unfolding fps_eq_iff fps_add_nth  fps_const_mult_left fps_deriv_nth by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1547
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1548
lemma fps_deriv_mult[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1549
  fixes f :: "'a::comm_ring_1 fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1550
  shows "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1551
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1552
  let ?D = "fps_deriv"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1553
  have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" for n
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1554
  proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1555
    let ?Zn = "{0 ..n}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1556
    let ?Zn1 = "{0 .. n + 1}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1557
    let ?g = "\<lambda>i. of_nat (i+1) * g $ (i+1) * f $ (n - i) +
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1558
        of_nat (i+1)* f $ (i+1) * g $ (n - i)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1559
    let ?h = "\<lambda>i. of_nat i * g $ i * f $ ((n+1) - i) +
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1560
        of_nat i* f $ i * g $ ((n + 1) - i)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1561
    have s0: "sum (\<lambda>i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 =
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1562
      sum (\<lambda>i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66817
diff changeset
  1563
       by (rule sum.reindex_bij_witness[where i="(-) (n + 1)" and j="(-) (n + 1)"]) auto
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1564
    have s1: "sum (\<lambda>i. f $ i * g $ (n + 1 - i)) ?Zn1 =
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1565
      sum (\<lambda>i. f $ (n + 1 - i) * g $ i) ?Zn1"
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66817
diff changeset
  1566
       by (rule sum.reindex_bij_witness[where i="(-) (n + 1)" and j="(-) (n + 1)"]) auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1567
    have "(f * ?D g + ?D f * g)$n = (?D g * f + ?D f * g)$n"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  1568
      by (simp only: mult.commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1569
    also have "\<dots> = (\<Sum>i = 0..n. ?g i)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1570
      by (simp add: fps_mult_nth sum.distrib[symmetric])
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1571
    also have "\<dots> = sum ?h {0..n+1}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1572
      by (rule sum.reindex_bij_witness_not_neutral
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  1573
            [where S'="{}" and T'="{0}" and j="Suc" and i="\<lambda>i. i - 1"]) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1574
    also have "\<dots> = (fps_deriv (f * g)) $ n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1575
      apply (simp only: fps_deriv_nth fps_mult_nth sum.distrib)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1576
      unfolding s0 s1
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1577
      unfolding sum.distrib[symmetric] sum_distrib_left
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1578
      apply (rule sum.cong)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1579
      apply (auto simp add: of_nat_diff field_simps)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1580
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1581
    finally show ?thesis .
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1582
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1583
  then show ?thesis
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1584
    unfolding fps_eq_iff by auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1585
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1586
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1587
lemma fps_deriv_fps_X[simp]: "fps_deriv fps_X = 1"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1588
  by (simp add: fps_deriv_def fps_X_def fps_eq_iff)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  1589
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1590
lemma fps_deriv_neg[simp]:
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1591
  "fps_deriv (- (f:: 'a::comm_ring_1 fps)) = - (fps_deriv f)"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  1592
  by (simp add: fps_eq_iff fps_deriv_def)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1593
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1594
lemma fps_deriv_add[simp]:
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1595
  "fps_deriv ((f:: 'a::comm_ring_1 fps) + g) = fps_deriv f + fps_deriv g"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1596
  using fps_deriv_linear[of 1 f 1 g] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1597
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1598
lemma fps_deriv_sub[simp]:
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1599
  "fps_deriv ((f:: 'a::comm_ring_1 fps) - g) = fps_deriv f - fps_deriv g"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  1600
  using fps_deriv_add [of f "- g"] by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1601
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1602
lemma fps_deriv_const[simp]: "fps_deriv (fps_const c) = 0"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  1603
  by (simp add: fps_ext fps_deriv_def fps_const_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1604
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  1605
lemma fps_deriv_of_nat [simp]: "fps_deriv (of_nat n) = 0"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  1606
  by (simp add: fps_of_nat [symmetric])
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  1607
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  1608
lemma fps_deriv_numeral [simp]: "fps_deriv (numeral n) = 0"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  1609
  by (simp add: numeral_fps_const)    
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  1610
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1611
lemma fps_deriv_mult_const_left[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1612
  "fps_deriv (fps_const (c::'a::comm_ring_1) * f) = fps_const c * fps_deriv f"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1613
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1614
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1615
lemma fps_deriv_0[simp]: "fps_deriv 0 = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1616
  by (simp add: fps_deriv_def fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1617
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1618
lemma fps_deriv_1[simp]: "fps_deriv 1 = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1619
  by (simp add: fps_deriv_def fps_eq_iff )
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1620
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1621
lemma fps_deriv_mult_const_right[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1622
  "fps_deriv (f * fps_const (c::'a::comm_ring_1)) = fps_deriv f * fps_const c"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1623
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1624
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1625
lemma fps_deriv_sum:
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1626
  "fps_deriv (sum f S) = sum (\<lambda>i. fps_deriv (f i :: 'a::comm_ring_1 fps)) S"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1627
proof (cases "finite S")
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1628
  case False
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1629
  then show ?thesis by simp
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1630
next
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1631
  case True
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1632
  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1633
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1634
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1635
lemma fps_deriv_eq_0_iff [simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1636
  "fps_deriv f = 0 \<longleftrightarrow> f = fps_const (f$0 :: 'a::{idom,semiring_char_0})"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1637
  (is "?lhs \<longleftrightarrow> ?rhs")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1638
proof
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1639
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1640
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1641
    from that have "fps_deriv f = fps_deriv (fps_const (f$0))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1642
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1643
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1644
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1645
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1646
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1647
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1648
    from that have "\<forall>n. (fps_deriv f)$n = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1649
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1650
    then have "\<forall>n. f$(n+1) = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1651
      by (simp del: of_nat_Suc of_nat_add One_nat_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1652
    then show ?thesis
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1653
      apply (clarsimp simp add: fps_eq_iff fps_const_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1654
      apply (erule_tac x="n - 1" in allE)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1655
      apply simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1656
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1657
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1658
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1659
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1660
lemma fps_deriv_eq_iff:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1661
  fixes f :: "'a::{idom,semiring_char_0} fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1662
  shows "fps_deriv f = fps_deriv g \<longleftrightarrow> (f = fps_const(f$0 - g$0) + g)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1663
proof -
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1664
  have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1665
    by simp
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1666
  also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f - g) $ 0)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1667
    unfolding fps_deriv_eq_0_iff ..
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1668
  finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1669
    by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1670
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1671
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1672
lemma fps_deriv_eq_iff_ex:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1673
  "(fps_deriv f = fps_deriv g) \<longleftrightarrow> (\<exists>c::'a::{idom,semiring_char_0}. f = fps_const c + g)"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1674
  by (auto simp: fps_deriv_eq_iff)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1675
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1676
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1677
fun fps_nth_deriv :: "nat \<Rightarrow> 'a::semiring_1 fps \<Rightarrow> 'a fps"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1678
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1679
  "fps_nth_deriv 0 f = f"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1680
| "fps_nth_deriv (Suc n) f = fps_nth_deriv n (fps_deriv f)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1681
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1682
lemma fps_nth_deriv_commute: "fps_nth_deriv (Suc n) f = fps_deriv (fps_nth_deriv n f)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1683
  by (induct n arbitrary: f) auto
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1684
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1685
lemma fps_nth_deriv_linear[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1686
  "fps_nth_deriv n (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1687
    fps_const a * fps_nth_deriv n f + fps_const b * fps_nth_deriv n g"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1688
  by (induct n arbitrary: f g) (auto simp add: fps_nth_deriv_commute)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1689
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1690
lemma fps_nth_deriv_neg[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1691
  "fps_nth_deriv n (- (f :: 'a::comm_ring_1 fps)) = - (fps_nth_deriv n f)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1692
  by (induct n arbitrary: f) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1693
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1694
lemma fps_nth_deriv_add[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1695
  "fps_nth_deriv n ((f :: 'a::comm_ring_1 fps) + g) = fps_nth_deriv n f + fps_nth_deriv n g"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1696
  using fps_nth_deriv_linear[of n 1 f 1 g] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1697
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1698
lemma fps_nth_deriv_sub[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1699
  "fps_nth_deriv n ((f :: 'a::comm_ring_1 fps) - g) = fps_nth_deriv n f - fps_nth_deriv n g"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  1700
  using fps_nth_deriv_add [of n f "- g"] by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1701
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1702
lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1703
  by (induct n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1704
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1705
lemma fps_nth_deriv_1[simp]: "fps_nth_deriv n 1 = (if n = 0 then 1 else 0)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1706
  by (induct n) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1707
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1708
lemma fps_nth_deriv_const[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1709
  "fps_nth_deriv n (fps_const c) = (if n = 0 then fps_const c else 0)"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1710
  by (cases n) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1711
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1712
lemma fps_nth_deriv_mult_const_left[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1713
  "fps_nth_deriv n (fps_const (c::'a::comm_ring_1) * f) = fps_const c * fps_nth_deriv n f"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1714
  using fps_nth_deriv_linear[of n "c" f 0 0 ] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1715
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1716
lemma fps_nth_deriv_mult_const_right[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1717
  "fps_nth_deriv n (f * fps_const (c::'a::comm_ring_1)) = fps_nth_deriv n f * fps_const c"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  1718
  using fps_nth_deriv_linear[of n "c" f 0 0] by (simp add: mult.commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1719
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1720
lemma fps_nth_deriv_sum:
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1721
  "fps_nth_deriv n (sum f S) = sum (\<lambda>i. fps_nth_deriv n (f i :: 'a::comm_ring_1 fps)) S"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1722
proof (cases "finite S")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1723
  case True
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1724
  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1725
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1726
  case False
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1727
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1728
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1729
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1730
lemma fps_deriv_maclauren_0:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1731
  "(fps_nth_deriv k (f :: 'a::comm_semiring_1 fps)) $ 0 = of_nat (fact k) * f $ k"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  1732
  by (induct k arbitrary: f) (auto simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1733
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1734
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1735
subsection \<open>Powers\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1736
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1737
lemma fps_power_zeroth_eq_one: "a$0 =1 \<Longrightarrow> a^n $ 0 = (1::'a::semiring_1)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1738
  by (induct n) (auto simp add: expand_fps_eq fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1739
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1740
lemma fps_power_first_eq: "(a :: 'a::comm_ring_1 fps) $ 0 =1 \<Longrightarrow> a^n $ 1 = of_nat n * a$1"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1741
proof (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1742
  case 0
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1743
  then show ?case by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1744
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1745
  case (Suc n)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1746
  show ?case unfolding power_Suc fps_mult_nth
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1747
    using Suc.hyps[OF \<open>a$0 = 1\<close>] \<open>a$0 = 1\<close> fps_power_zeroth_eq_one[OF \<open>a$0=1\<close>]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1748
    by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1749
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1750
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1751
lemma startsby_one_power:"a $ 0 = (1::'a::comm_ring_1) \<Longrightarrow> a^n $ 0 = 1"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1752
  by (induct n) (auto simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1753
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1754
lemma startsby_zero_power:"a $0 = (0::'a::comm_ring_1) \<Longrightarrow> n > 0 \<Longrightarrow> a^n $0 = 0"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1755
  by (induct n) (auto simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1756
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1757
lemma startsby_power:"a $0 = (v::'a::comm_ring_1) \<Longrightarrow> a^n $0 = v^n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1758
  by (induct n) (auto simp add: fps_mult_nth)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1759
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1760
lemma startsby_zero_power_iff[simp]: "a^n $0 = (0::'a::idom) \<longleftrightarrow> n \<noteq> 0 \<and> a$0 = 0"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1761
  apply (rule iffI)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1762
  apply (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1763
  apply (auto simp add: fps_mult_nth)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1764
  apply (rule startsby_zero_power, simp_all)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1765
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1766
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1767
lemma startsby_zero_power_prefix:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1768
  assumes a0: "a $ 0 = (0::'a::idom)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1769
  shows "\<forall>n < k. a ^ k $ n = 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1770
  using a0
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1771
proof (induct k rule: nat_less_induct)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1772
  fix k
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1773
  assume H: "\<forall>m<k. a $0 =  0 \<longrightarrow> (\<forall>n<m. a ^ m $ n = 0)" and a0: "a $ 0 = 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1774
  show "\<forall>m<k. a ^ k $ m = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1775
  proof (cases k)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1776
    case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1777
    then show ?thesis by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1778
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1779
    case (Suc l)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1780
    have "a^k $ m = 0" if mk: "m < k" for m
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1781
    proof (cases "m = 0")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1782
      case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1783
      then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1784
        using startsby_zero_power[of a k] Suc a0 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1785
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1786
      case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1787
      have "a ^k $ m = (a^l * a) $m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1788
        by (simp add: Suc mult.commute)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1789
      also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1790
        by (simp add: fps_mult_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1791
      also have "\<dots> = 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1792
        apply (rule sum.neutral)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1793
        apply auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1794
        apply (case_tac "x = m")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1795
        using a0 apply simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1796
        apply (rule H[rule_format])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1797
        using a0 Suc mk apply auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1798
        done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1799
      finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1800
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1801
    then show ?thesis by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1802
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1803
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1804
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1805
lemma startsby_zero_sum_depends:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1806
  assumes a0: "a $0 = (0::'a::idom)"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1807
    and kn: "n \<ge> k"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1808
  shows "sum (\<lambda>i. (a ^ i)$k) {0 .. n} = sum (\<lambda>i. (a ^ i)$k) {0 .. k}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1809
  apply (rule sum.mono_neutral_right)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1810
  using kn
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1811
  apply auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1812
  apply (rule startsby_zero_power_prefix[rule_format, OF a0])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1813
  apply arith
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1814
  done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1815
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1816
lemma startsby_zero_power_nth_same:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1817
  assumes a0: "a$0 = (0::'a::idom)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1818
  shows "a^n $ n = (a$1) ^ n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1819
proof (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1820
  case 0
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1821
  then show ?case by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1822
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1823
  case (Suc n)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1824
  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1825
    by (simp add: field_simps)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1826
  also have "\<dots> = sum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1827
    by (simp add: fps_mult_nth)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1828
  also have "\<dots> = sum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1829
    apply (rule sum.mono_neutral_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1830
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1831
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1832
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1833
    apply (rule startsby_zero_power_prefix[rule_format, OF a0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1834
    apply arith
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1835
    done
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1836
  also have "\<dots> = a^n $ n * a$1"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1837
    using a0 by simp
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1838
  finally show ?case
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1839
    using Suc.hyps by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1840
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1841
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1842
lemma fps_inverse_power:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1843
  fixes a :: "'a::field fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1844
  shows "inverse (a^n) = inverse a ^ n"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1845
  by (induction n) (simp_all add: fps_inverse_mult)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1846
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1847
lemma fps_deriv_power:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1848
  "fps_deriv (a ^ n) = fps_const (of_nat n :: 'a::comm_ring_1) * fps_deriv a * a ^ (n - 1)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1849
  apply (induct n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1850
  apply (auto simp add: field_simps fps_const_add[symmetric] simp del: fps_const_add)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1851
  apply (case_tac n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1852
  apply (auto simp add: field_simps)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1853
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1854
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1855
lemma fps_inverse_deriv:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1856
  fixes a :: "'a::field fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1857
  assumes a0: "a$0 \<noteq> 0"
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1858
  shows "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1859
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1860
  from inverse_mult_eq_1[OF a0]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1861
  have "fps_deriv (inverse a * a) = 0" by simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  1862
  then have "inverse a * fps_deriv a + fps_deriv (inverse a) * a = 0"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  1863
    by simp
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  1864
  then have "inverse a * (inverse a * fps_deriv a + fps_deriv (inverse a) * a) = 0"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  1865
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1866
  with inverse_mult_eq_1[OF a0]
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1867
  have "(inverse a)\<^sup>2 * fps_deriv a + fps_deriv (inverse a) = 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1868
    unfolding power2_eq_square
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1869
    apply (simp add: field_simps)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  1870
    apply (simp add: mult.assoc[symmetric])
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1871
    done
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1872
  then have "(inverse a)\<^sup>2 * fps_deriv a + fps_deriv (inverse a) - fps_deriv a * (inverse a)\<^sup>2 =
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1873
      0 - fps_deriv a * (inverse a)\<^sup>2"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1874
    by simp
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1875
  then show "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1876
    by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1877
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1878
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1879
lemma fps_inverse_deriv':
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1880
  fixes a :: "'a::field fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1881
  assumes a0: "a $ 0 \<noteq> 0"
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1882
  shows "fps_deriv (inverse a) = - fps_deriv a / a\<^sup>2"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1883
  using fps_inverse_deriv[OF a0] a0
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1884
  by (simp add: fps_divide_unit power2_eq_square fps_inverse_mult)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1885
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1886
lemma inverse_mult_eq_1':
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1887
  assumes f0: "f$0 \<noteq> (0::'a::field)"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  1888
  shows "f * inverse f = 1"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  1889
  by (metis mult.commute inverse_mult_eq_1 f0)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1890
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1891
lemma fps_inverse_minus [simp]: "inverse (-f) = -inverse (f :: 'a :: field fps)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1892
  by (cases "f$0 = 0") (auto intro: fps_inverse_unique simp: inverse_mult_eq_1' fps_inverse_eq_0)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1893
  
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1894
lemma divide_fps_const [simp]: "f / fps_const (c :: 'a :: field) = fps_const (inverse c) * f"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1895
  by (cases "c = 0") (simp_all add: fps_divide_unit fps_const_inverse)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1896
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1897
(* FIfps_XME: The last part of this proof should go through by simp once we have a proper
61804
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1898
   theorem collection for simplifying division on rings *)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1899
lemma fps_divide_deriv:
61804
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1900
  assumes "b dvd (a :: 'a :: field fps)"
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1901
  shows   "fps_deriv (a / b) = (fps_deriv a * b - a * fps_deriv b) / b^2"
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1902
proof -
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1903
  have eq_divide_imp: "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b div c" for a b c :: "'a :: field fps"
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1904
    by (drule sym) (simp add: mult.assoc)
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1905
  from assms have "a = a / b * b" by simp
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1906
  also have "fps_deriv (a / b * b) = fps_deriv (a / b) * b + a / b * fps_deriv b" by simp
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1907
  finally have "fps_deriv (a / b) * b^2 = fps_deriv a * b - a * fps_deriv b" using assms
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1908
    by (simp add: power2_eq_square algebra_simps)
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1909
  thus ?thesis by (cases "b = 0") (auto simp: eq_divide_imp)
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1910
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1911
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1912
lemma fps_inverse_gp': "inverse (Abs_fps (\<lambda>n. 1::'a::field)) = 1 - fps_X"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1913
  by (simp add: fps_inverse_gp fps_eq_iff fps_X_def)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1914
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1915
lemma fps_one_over_one_minus_fps_X_squared:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1916
  "inverse ((1 - fps_X)^2 :: 'a :: field fps) = Abs_fps (\<lambda>n. of_nat (n+1))"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1917
proof -
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1918
  have "inverse ((1 - fps_X)^2 :: 'a fps) = fps_deriv (inverse (1 - fps_X))"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1919
    by (subst fps_inverse_deriv) (simp_all add: fps_inverse_power)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1920
  also have "inverse (1 - fps_X :: 'a fps) = Abs_fps (\<lambda>_. 1)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1921
    by (subst fps_inverse_gp' [symmetric]) simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1922
  also have "fps_deriv \<dots> = Abs_fps (\<lambda>n. of_nat (n + 1))"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1923
    by (simp add: fps_deriv_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1924
  finally show ?thesis .
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1925
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  1926
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1927
lemma fps_nth_deriv_fps_X[simp]: "fps_nth_deriv n fps_X = (if n = 0 then fps_X else if n=1 then 1 else 0)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1928
  by (cases n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1929
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1930
lemma fps_inverse_fps_X_plus1: "inverse (1 + fps_X) = Abs_fps (\<lambda>n. (- (1::'a::field)) ^ n)"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1931
  (is "_ = ?r")
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1932
proof -
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1933
  have eq: "(1 + fps_X) * ?r = 1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1934
    unfolding minus_one_power_iff
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1935
    by (auto simp add: field_simps fps_eq_iff)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1936
  show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1937
    by (auto simp add: eq intro: fps_inverse_unique)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1938
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1939
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1940
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1941
subsection \<open>Integration\<close>
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1942
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1943
definition fps_integral :: "'a::field_char_0 fps \<Rightarrow> 'a \<Rightarrow> 'a fps"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1944
  where "fps_integral a a0 = Abs_fps (\<lambda>n. if n = 0 then a0 else (a$(n - 1) / of_nat n))"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1945
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1946
lemma fps_deriv_fps_integral: "fps_deriv (fps_integral a a0) = a"
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1947
  unfolding fps_integral_def fps_deriv_def
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1948
  by (simp add: fps_eq_iff del: of_nat_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1949
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1950
lemma fps_integral_linear:
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1951
  "fps_integral (fps_const a * f + fps_const b * g) (a*a0 + b*b0) =
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1952
    fps_const a * fps_integral f a0 + fps_const b * fps_integral g b0"
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1953
  (is "?l = ?r")
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1954
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1955
  have "fps_deriv ?l = fps_deriv ?r"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1956
    by (simp add: fps_deriv_fps_integral)
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1957
  moreover have "?l$0 = ?r$0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1958
    by (simp add: fps_integral_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1959
  ultimately show ?thesis
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1960
    unfolding fps_deriv_eq_iff by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1961
qed
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1962
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1963
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1964
subsection \<open>Composition of FPSs\<close>
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1965
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1966
definition fps_compose :: "'a::semiring_1 fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps"  (infixl "oo" 55)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1967
  where "a oo b = Abs_fps (\<lambda>n. sum (\<lambda>i. a$i * (b^i$n)) {0..n})"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1968
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1969
lemma fps_compose_nth: "(a oo b)$n = sum (\<lambda>i. a$i * (b^i$n)) {0..n}"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1970
  by (simp add: fps_compose_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1971
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1972
lemma fps_compose_nth_0 [simp]: "(f oo g) $ 0 = f $ 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1973
  by (simp add: fps_compose_nth)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1974
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1975
lemma fps_compose_fps_X[simp]: "a oo fps_X = (a :: 'a::comm_ring_1 fps)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1976
  by (simp add: fps_ext fps_compose_def mult_delta_right sum.delta')
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1977
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1978
lemma fps_const_compose[simp]: "fps_const (a::'a::comm_ring_1) oo b = fps_const a"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1979
  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1980
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1981
lemma numeral_compose[simp]: "(numeral k :: 'a::comm_ring_1 fps) oo b = numeral k"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1982
  unfolding numeral_fps_const by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1983
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1984
lemma neg_numeral_compose[simp]: "(- numeral k :: 'a::comm_ring_1 fps) oo b = - numeral k"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1985
  unfolding neg_numeral_fps_const by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  1986
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1987
lemma fps_X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> fps_X oo a = (a :: 'a::comm_ring_1 fps)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1988
  by (simp add: fps_eq_iff fps_compose_def mult_delta_left sum.delta not_le)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1989
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1990
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1991
subsection \<open>Rules from Herbert Wilf's Generatingfunctionology\<close>
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1992
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1993
subsubsection \<open>Rule 1\<close>
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  1994
  (* {a_{n+k}}_0^infty Corresponds to (f - sum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1995
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1996
lemma fps_power_mult_eq_shift:
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1997
  "fps_X^Suc k * Abs_fps (\<lambda>n. a (n + Suc k)) =
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  1998
    Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a::comm_ring_1) * fps_X^i) {0 .. k}"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1999
  (is "?lhs = ?rhs")
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2000
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2001
  have "?lhs $ n = ?rhs $ n" for n :: nat
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2002
  proof -
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2003
    have "?lhs $ n = (if n < Suc k then 0 else a n)"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2004
      unfolding fps_X_power_mult_nth by auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2005
    also have "\<dots> = ?rhs $ n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2006
    proof (induct k)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2007
      case 0
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2008
      then show ?case
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2009
        by (simp add: fps_sum_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2010
    next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2011
      case (Suc k)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2012
      have "(Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a) * fps_X^i) {0 .. Suc k})$n =
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2013
        (Abs_fps a - sum (\<lambda>i. fps_const (a i :: 'a) * fps_X^i) {0 .. k} -
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2014
          fps_const (a (Suc k)) * fps_X^ Suc k) $ n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2015
        by (simp add: field_simps)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2016
      also have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * fps_X^ Suc k)$n"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2017
        using Suc.hyps[symmetric] unfolding fps_sub_nth by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2018
      also have "\<dots> = (if n < Suc (Suc k) then 0 else a n)"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2019
        unfolding fps_X_power_mult_right_nth
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2020
        apply (auto simp add: not_less fps_const_def)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2021
        apply (rule cong[of a a, OF refl])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2022
        apply arith
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2023
        done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2024
      finally show ?case
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2025
        by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2026
    qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2027
    finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2028
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2029
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2030
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2031
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2032
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2033
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2034
subsubsection \<open>Rule 2\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2035
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2036
  (* We can not reach the form of Wilf, but still near to it using rewrite rules*)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2037
  (* If f reprents {a_n} and P is a polynomial, then
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2038
        P(xD) f represents {P(n) a_n}*)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2039
69064
5840724b1d71 Prefix form of infix with * on either side no longer needs special treatment
nipkow
parents: 68975
diff changeset
  2040
definition "fps_XD = (*) fps_X \<circ> fps_deriv"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2041
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2042
lemma fps_XD_add[simp]:"fps_XD (a + b) = fps_XD a + fps_XD (b :: 'a::comm_ring_1 fps)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2043
  by (simp add: fps_XD_def field_simps)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2044
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2045
lemma fps_XD_mult_const[simp]:"fps_XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * fps_XD a"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2046
  by (simp add: fps_XD_def field_simps)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2047
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2048
lemma fps_XD_linear[simp]: "fps_XD (fps_const c * a + fps_const d * b) =
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2049
    fps_const c * fps_XD a + fps_const d * fps_XD (b :: 'a::comm_ring_1 fps)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2050
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2051
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2052
lemma fps_XDN_linear:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2053
  "(fps_XD ^^ n) (fps_const c * a + fps_const d * b) =
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2054
    fps_const c * (fps_XD ^^ n) a + fps_const d * (fps_XD ^^ n) (b :: 'a::comm_ring_1 fps)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2055
  by (induct n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2056
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2057
lemma fps_mult_fps_X_deriv_shift: "fps_X* fps_deriv a = Abs_fps (\<lambda>n. of_nat n* a$n)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2058
  by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2059
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2060
lemma fps_mult_fps_XD_shift:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2061
  "(fps_XD ^^ k) (a :: 'a::comm_ring_1 fps) = Abs_fps (\<lambda>n. (of_nat n ^ k) * a$n)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2062
  by (induct k arbitrary: a) (simp_all add: fps_XD_def fps_eq_iff field_simps del: One_nat_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2063
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2064
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2065
subsubsection \<open>Rule 3\<close>
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2066
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61552
diff changeset
  2067
text \<open>Rule 3 is trivial and is given by \<open>fps_times_def\<close>.\<close>
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2068
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2069
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2070
subsubsection \<open>Rule 5 --- summation and "division" by (1 - fps_X)\<close>
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2071
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2072
lemma fps_divide_fps_X_minus1_sum_lemma:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2073
  "a = ((1::'a::comm_ring_1 fps) - fps_X) * Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2074
proof -
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2075
  let ?sa = "Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2076
  have th0: "\<And>i. (1 - (fps_X::'a fps)) $ i = (if i = 0 then 1 else if i = 1 then - 1 else 0)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2077
    by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2078
  have "a$n = ((1 - fps_X) * ?sa) $ n" for n
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2079
  proof (cases "n = 0")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2080
    case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2081
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2082
      by (simp add: fps_mult_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2083
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2084
    case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2085
    then have u: "{0} \<union> ({1} \<union> {2..n}) = {0..n}" "{1} \<union> {2..n} = {1..n}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2086
      "{0..n - 1} \<union> {n} = {0..n}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2087
      by (auto simp: set_eq_iff)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2088
    have d: "{0} \<inter> ({1} \<union> {2..n}) = {}" "{1} \<inter> {2..n} = {}" "{0..n - 1} \<inter> {n} = {}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2089
      using False by simp_all
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2090
    have f: "finite {0}" "finite {1}" "finite {2 .. n}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2091
      "finite {0 .. n - 1}" "finite {n}" by simp_all
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2092
    have "((1 - fps_X) * ?sa) $ n = sum (\<lambda>i. (1 - fps_X)$ i * ?sa $ (n - i)) {0 .. n}"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2093
      by (simp add: fps_mult_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2094
    also have "\<dots> = a$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2095
      unfolding th0
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2096
      unfolding sum.union_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2097
      unfolding sum.union_disjoint[OF f(2) f(3) d(2)]
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2098
      apply (simp)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2099
      unfolding sum.union_disjoint[OF f(4,5) d(3), unfolded u(3)]
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2100
      apply simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2101
      done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2102
    finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2103
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2104
  qed
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2105
  then show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2106
    unfolding fps_eq_iff by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2107
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2108
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2109
lemma fps_divide_fps_X_minus1_sum:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2110
  "a /((1::'a::field fps) - fps_X) = Abs_fps (\<lambda>n. sum (\<lambda>i. a $ i) {0..n})"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2111
proof -
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2112
  let ?fps_X = "1 - (fps_X::'a fps)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2113
  have th0: "?fps_X $ 0 \<noteq> 0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2114
    by simp
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66817
diff changeset
  2115
  have "a /?fps_X = ?fps_X *  Abs_fps (\<lambda>n::nat. sum (($) a) {0..n}) * inverse ?fps_X"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2116
    using fps_divide_fps_X_minus1_sum_lemma[of a, symmetric] th0
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  2117
    by (simp add: fps_divide_def mult.assoc)
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66817
diff changeset
  2118
  also have "\<dots> = (inverse ?fps_X * ?fps_X) * Abs_fps (\<lambda>n::nat. sum (($) a) {0..n}) "
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  2119
    by (simp add: ac_simps)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2120
  finally show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2121
    by (simp add: inverse_mult_eq_1[OF th0])
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2122
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2123
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2124
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2125
subsubsection \<open>Rule 4 in its more general form: generalizes Rule 3 for an arbitrary
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2126
  finite product of FPS, also the relvant instance of powers of a FPS\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2127
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2128
definition "natpermute n k = {l :: nat list. length l = k \<and> sum_list l = n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2129
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2130
lemma natlist_trivial_1: "natpermute n 1 = {[n]}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2131
  apply (auto simp add: natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2132
  apply (case_tac x)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2133
  apply auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2134
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2135
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2136
lemma append_natpermute_less_eq:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2137
  assumes "xs @ ys \<in> natpermute n k"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2138
  shows "sum_list xs \<le> n"
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2139
    and "sum_list ys \<le> n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2140
proof -
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2141
  from assms have "sum_list (xs @ ys) = n"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2142
    by (simp add: natpermute_def)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2143
  then have "sum_list xs + sum_list ys = n"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2144
    by simp
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2145
  then show "sum_list xs \<le> n" and "sum_list ys \<le> n"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2146
    by simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2147
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2148
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2149
lemma natpermute_split:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2150
  assumes "h \<le> k"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2151
  shows "natpermute n k =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2152
    (\<Union>m \<in>{0..n}. {l1 @ l2 |l1 l2. l1 \<in> natpermute m h \<and> l2 \<in> natpermute (n - m) (k - h)})"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2153
  (is "?L = ?R" is "_ = (\<Union>m \<in>{0..n}. ?S m)")
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2154
proof
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2155
  show "?R \<subseteq> ?L"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2156
  proof
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2157
    fix l
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2158
    assume l: "l \<in> ?R"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2159
    from l obtain m xs ys where h: "m \<in> {0..n}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2160
      and xs: "xs \<in> natpermute m h"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2161
      and ys: "ys \<in> natpermute (n - m) (k - h)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2162
      and leq: "l = xs@ys" by blast
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2163
    from xs have xs': "sum_list xs = m"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2164
      by (simp add: natpermute_def)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2165
    from ys have ys': "sum_list ys = n - m"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2166
      by (simp add: natpermute_def)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2167
    show "l \<in> ?L" using leq xs ys h
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2168
      apply (clarsimp simp add: natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2169
      unfolding xs' ys'
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2170
      using assms xs ys
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2171
      unfolding natpermute_def
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2172
      apply simp
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2173
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2174
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2175
  show "?L \<subseteq> ?R"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2176
  proof
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2177
    fix l
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2178
    assume l: "l \<in> natpermute n k"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2179
    let ?xs = "take h l"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2180
    let ?ys = "drop h l"
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2181
    let ?m = "sum_list ?xs"
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2182
    from l have ls: "sum_list (?xs @ ?ys) = n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2183
      by (simp add: natpermute_def)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2184
    have xs: "?xs \<in> natpermute ?m h" using l assms
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2185
      by (simp add: natpermute_def)
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2186
    have l_take_drop: "sum_list l = sum_list (take h l @ drop h l)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2187
      by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2188
    then have ys: "?ys \<in> natpermute (n - ?m) (k - h)"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2189
      using l assms ls by (auto simp add: natpermute_def simp del: append_take_drop_id)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2190
    from ls have m: "?m \<in> {0..n}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2191
      by (simp add: l_take_drop del: append_take_drop_id)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2192
    from xs ys ls show "l \<in> ?R"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2193
      apply auto
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2194
      apply (rule bexI [where x = "?m"])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2195
      apply (rule exI [where x = "?xs"])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2196
      apply (rule exI [where x = "?ys"])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2197
      using ls l
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2198
      apply (auto simp add: natpermute_def l_take_drop simp del: append_take_drop_id)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2199
      apply simp
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2200
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2201
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2202
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2203
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2204
lemma natpermute_0: "natpermute n 0 = (if n = 0 then {[]} else {})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2205
  by (auto simp add: natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2206
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2207
lemma natpermute_0'[simp]: "natpermute 0 k = (if k = 0 then {[]} else {replicate k 0})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2208
  apply (auto simp add: set_replicate_conv_if natpermute_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2209
  apply (rule nth_equalityI)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2210
  apply simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2211
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2212
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2213
lemma natpermute_finite: "finite (natpermute n k)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2214
proof (induct k arbitrary: n)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2215
  case 0
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2216
  then show ?case
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2217
    apply (subst natpermute_split[of 0 0, simplified])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2218
    apply (simp add: natpermute_0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2219
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2220
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2221
  case (Suc k)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2222
  then show ?case unfolding natpermute_split [of k "Suc k", simplified]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2223
    apply -
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2224
    apply (rule finite_UN_I)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2225
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2226
    unfolding One_nat_def[symmetric] natlist_trivial_1
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2227
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2228
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2229
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2230
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2231
lemma natpermute_contain_maximal:
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2232
  "{xs \<in> natpermute n (k + 1). n \<in> set xs} = (\<Union>i\<in>{0 .. k}. {(replicate (k + 1) 0) [i:=n]})"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2233
  (is "?A = ?B")
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2234
proof
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2235
  show "?A \<subseteq> ?B"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2236
  proof
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2237
    fix xs
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2238
    assume "xs \<in> ?A"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2239
    then have H: "xs \<in> natpermute n (k + 1)" and n: "n \<in> set xs"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2240
      by blast+
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2241
    then obtain i where i: "i \<in> {0.. k}" "xs!i = n"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2242
      unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2243
    have eqs: "({0..k} - {i}) \<union> {i} = {0..k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2244
      using i by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2245
    have f: "finite({0..k} - {i})" "finite {i}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2246
      by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2247
    have d: "({0..k} - {i}) \<inter> {i} = {}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2248
      using i by auto
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2249
    from H have "n = sum (nth xs) {0..k}"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2250
      apply (simp add: natpermute_def)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2251
      apply (auto simp add: atLeastLessThanSuc_atLeastAtMost sum_list_sum_nth)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2252
      done
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2253
    also have "\<dots> = n + sum (nth xs) ({0..k} - {i})"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2254
      unfolding sum.union_disjoint[OF f d, unfolded eqs] using i by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2255
    finally have zxs: "\<forall> j\<in> {0..k} - {i}. xs!j = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2256
      by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2257
    from H have xsl: "length xs = k+1"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2258
      by (simp add: natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2259
    from i have i': "i < length (replicate (k+1) 0)"   "i < k+1"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2260
      unfolding length_replicate by presburger+
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2261
    have "xs = (replicate (k+1) 0) [i := n]"
68975
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2262
    proof (rule nth_equalityI)
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2263
      show "length xs = length ((replicate (k + 1) 0)[i := n])"
68975
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2264
        by (metis length_list_update length_replicate xsl)
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2265
      show "xs ! j = (replicate (k + 1) 0)[i := n] ! j" if "j < length xs" for j
68975
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2266
      proof (cases "j = i")
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2267
        case True
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2268
        then show ?thesis
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2269
          by (metis i'(1) i(2) nth_list_update)
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2270
      next
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2271
        case False
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2272
        with that show ?thesis
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2273
          by (simp add: xsl zxs del: replicate.simps split: nat.split)
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2274
      qed
5ce4d117cea7 A few new results, elimination of duplicates and more use of "pairwise"
paulson <lp15@cam.ac.uk>
parents: 68442
diff changeset
  2275
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2276
    then show "xs \<in> ?B" using i by blast
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2277
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2278
  show "?B \<subseteq> ?A"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2279
  proof
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2280
    fix xs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2281
    assume "xs \<in> ?B"
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2282
    then obtain i where i: "i \<in> {0..k}" and xs: "xs = (replicate (k + 1) 0) [i:=n]"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2283
      by auto
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2284
    have nxs: "n \<in> set xs"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2285
      unfolding xs
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2286
      apply (rule set_update_memI)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2287
      using i apply simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2288
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2289
    have xsl: "length xs = k + 1"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2290
      by (simp only: xs length_replicate length_list_update)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2291
    have "sum_list xs = sum (nth xs) {0..<k+1}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2292
      unfolding sum_list_sum_nth xsl ..
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2293
    also have "\<dots> = sum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2294
      by (rule sum.cong) (simp_all add: xs del: replicate.simps)
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2295
    also have "\<dots> = n" using i by (simp)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2296
    finally have "xs \<in> natpermute n (k + 1)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2297
      using xsl unfolding natpermute_def mem_Collect_eq by blast
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2298
    then show "xs \<in> ?A"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2299
      using nxs by blast
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2300
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2301
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2302
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2303
text \<open>The general form.\<close>
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2304
lemma fps_prod_nth:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2305
  fixes m :: nat
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2306
    and a :: "nat \<Rightarrow> 'a::comm_ring_1 fps"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2307
  shows "(prod a {0 .. m}) $ n =
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2308
    sum (\<lambda>v. prod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2309
  (is "?P m n")
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2310
proof (induct m arbitrary: n rule: nat_less_induct)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2311
  fix m n assume H: "\<forall>m' < m. \<forall>n. ?P m' n"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2312
  show "?P m n"
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2313
  proof (cases m)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2314
    case 0
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2315
    then show ?thesis
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2316
      apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2317
      unfolding natlist_trivial_1[where n = n, unfolded One_nat_def]
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2318
      apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2319
      done
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2320
  next
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2321
    case (Suc k)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2322
    then have km: "k < m" by arith
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2323
    have u0: "{0 .. k} \<union> {m} = {0..m}"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2324
      using Suc by (simp add: set_eq_iff) presburger
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2325
    have f0: "finite {0 .. k}" "finite {m}" by auto
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2326
    have d0: "{0 .. k} \<inter> {m} = {}" using Suc by auto
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2327
    have "(prod a {0 .. m}) $ n = (prod a {0 .. k} * a m) $ n"
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2328
      unfolding prod.union_disjoint[OF f0 d0, unfolded u0] by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2329
    also have "\<dots> = (\<Sum>i = 0..n. (\<Sum>v\<in>natpermute i (k + 1). \<Prod>j\<in>{0..k}. a j $ v ! j) * a m $ (n - i))"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2330
      unfolding fps_mult_nth H[rule_format, OF km] ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2331
    also have "\<dots> = (\<Sum>v\<in>natpermute n (m + 1). \<Prod>j\<in>{0..m}. a j $ v ! j)"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2332
      apply (simp add: Suc)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2333
      unfolding natpermute_split[of m "m + 1", simplified, of n,
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2334
        unfolded natlist_trivial_1[unfolded One_nat_def] Suc]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2335
      apply (subst sum.UNION_disjoint)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2336
      apply simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2337
      apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2338
      unfolding image_Collect[symmetric]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2339
      apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2340
      apply (rule finite_imageI)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2341
      apply (rule natpermute_finite)
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  2342
      apply (clarsimp simp add: set_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2343
      apply auto
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2344
      apply (rule sum.cong)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2345
      apply (rule refl)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2346
      unfolding sum_distrib_right
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2347
      apply (rule sym)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2348
      apply (rule_tac l = "\<lambda>xs. xs @ [n - x]" in sum.reindex_cong)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2349
      apply (simp add: inj_on_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2350
      apply auto
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2351
      unfolding prod.union_disjoint[OF f0 d0, unfolded u0, unfolded Suc]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2352
      apply (clarsimp simp add: natpermute_def nth_append)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2353
      done
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2354
    finally show ?thesis .
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2355
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2356
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2357
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2358
text \<open>The special form for powers.\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2359
lemma fps_power_nth_Suc:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2360
  fixes m :: nat
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2361
    and a :: "'a::comm_ring_1 fps"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2362
  shows "(a ^ Suc m)$n = sum (\<lambda>v. prod (\<lambda>j. a $ (v!j)) {0..m}) (natpermute n (m+1))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2363
proof -
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2364
  have th0: "a^Suc m = prod (\<lambda>i. a) {0..m}"
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2365
    by (simp add: prod_constant)
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2366
  show ?thesis unfolding th0 fps_prod_nth ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2367
qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2368
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2369
lemma fps_power_nth:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2370
  fixes m :: nat
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2371
    and a :: "'a::comm_ring_1 fps"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2372
  shows "(a ^m)$n =
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2373
    (if m=0 then 1$n else sum (\<lambda>v. prod (\<lambda>j. a $ (v!j)) {0..m - 1}) (natpermute n m))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2374
  by (cases m) (simp_all add: fps_power_nth_Suc del: power_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2375
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2376
lemma fps_nth_power_0:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2377
  fixes m :: nat
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2378
    and a :: "'a::comm_ring_1 fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2379
  shows "(a ^m)$0 = (a$0) ^ m"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2380
proof (cases m)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2381
  case 0
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2382
  then show ?thesis by simp
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2383
next
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2384
  case (Suc n)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2385
  then have c: "m = card {0..n}" by simp
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2386
  have "(a ^m)$0 = prod (\<lambda>i. a$0) {0..n}"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2387
    by (simp add: Suc fps_power_nth del: replicate.simps power_Suc)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2388
  also have "\<dots> = (a$0) ^ m"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2389
   unfolding c by (rule prod_constant)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2390
 finally show ?thesis .
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2391
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2392
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2393
lemma natpermute_max_card:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2394
  assumes n0: "n \<noteq> 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2395
  shows "card {xs \<in> natpermute n (k + 1). n \<in> set xs} = k + 1"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2396
  unfolding natpermute_contain_maximal
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2397
proof -
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2398
  let ?A = "\<lambda>i. {(replicate (k + 1) 0)[i := n]}"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2399
  let ?K = "{0 ..k}"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2400
  have fK: "finite ?K"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2401
    by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2402
  have fAK: "\<forall>i\<in>?K. finite (?A i)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2403
    by auto
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2404
  have d: "\<forall>i\<in> ?K. \<forall>j\<in> ?K. i \<noteq> j \<longrightarrow>
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2405
    {(replicate (k + 1) 0)[i := n]} \<inter> {(replicate (k + 1) 0)[j := n]} = {}"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2406
  proof clarify
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2407
    fix i j
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2408
    assume i: "i \<in> ?K" and j: "j \<in> ?K" and ij: "i \<noteq> j"
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2409
    have False if eq: "(replicate (k+1) 0)[i:=n] = (replicate (k+1) 0)[j:= n]"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2410
    proof -
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2411
      have "(replicate (k+1) 0) [i:=n] ! i = n"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2412
        using i by (simp del: replicate.simps)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2413
      moreover
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2414
      have "(replicate (k+1) 0) [j:=n] ! i = 0"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2415
        using i ij by (simp del: replicate.simps)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2416
      ultimately show ?thesis
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2417
        using eq n0 by (simp del: replicate.simps)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2418
    qed
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2419
    then show "{(replicate (k + 1) 0)[i := n]} \<inter> {(replicate (k + 1) 0)[j := n]} = {}"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2420
      by auto
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2421
  qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2422
  from card_UN_disjoint[OF fK fAK d]
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2423
  show "card (\<Union>i\<in>{0..k}. {(replicate (k + 1) 0)[i := n]}) = k + 1"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2424
    by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2425
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2426
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2427
lemma fps_power_Suc_nth:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2428
  fixes f :: "'a :: comm_ring_1 fps"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2429
  assumes k: "k > 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2430
  shows "(f ^ Suc m) $ k = 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2431
           of_nat (Suc m) * (f $ k * (f $ 0) ^ m) +
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2432
           (\<Sum>v\<in>{v\<in>natpermute k (m+1). k \<notin> set v}. \<Prod>j = 0..m. f $ v ! j)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2433
proof -
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2434
  define A B 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2435
    where "A = {v\<in>natpermute k (m+1). k \<in> set v}" 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2436
      and  "B = {v\<in>natpermute k (m+1). k \<notin> set v}"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2437
  have [simp]: "finite A" "finite B" "A \<inter> B = {}" by (auto simp: A_def B_def natpermute_finite)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2438
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2439
  from natpermute_max_card[of k m] k have card_A: "card A = m + 1" by (simp add: A_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2440
  {
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2441
    fix v assume v: "v \<in> A"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2442
    from v have [simp]: "length v = Suc m" by (simp add: A_def natpermute_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2443
    from v have "\<exists>j. j \<le> m \<and> v ! j = k" 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2444
      by (auto simp: set_conv_nth A_def natpermute_def less_Suc_eq_le)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2445
    then guess j by (elim exE conjE) note j = this
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2446
    
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2447
    from v have "k = sum_list v" by (simp add: A_def natpermute_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2448
    also have "\<dots> = (\<Sum>i=0..m. v ! i)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2449
      by (simp add: sum_list_sum_nth atLeastLessThanSuc_atLeastAtMost del: sum_op_ivl_Suc)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2450
    also from j have "{0..m} = insert j ({0..m}-{j})" by auto
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2451
    also from j have "(\<Sum>i\<in>\<dots>. v ! i) = k + (\<Sum>i\<in>{0..m}-{j}. v ! i)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2452
      by (subst sum.insert) simp_all
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2453
    finally have "(\<Sum>i\<in>{0..m}-{j}. v ! i) = 0" by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2454
    hence zero: "v ! i = 0" if "i \<in> {0..m}-{j}" for i using that
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2455
      by (subst (asm) sum_eq_0_iff) auto
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2456
      
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2457
    from j have "{0..m} = insert j ({0..m} - {j})" by auto
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2458
    also from j have "(\<Prod>i\<in>\<dots>. f $ (v ! i)) = f $ k * (\<Prod>i\<in>{0..m} - {j}. f $ (v ! i))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2459
      by (subst prod.insert) auto
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2460
    also have "(\<Prod>i\<in>{0..m} - {j}. f $ (v ! i)) = (\<Prod>i\<in>{0..m} - {j}. f $ 0)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2461
      by (intro prod.cong) (simp_all add: zero)
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2462
    also from j have "\<dots> = (f $ 0) ^ m" by (subst prod_constant) simp_all
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2463
    finally have "(\<Prod>j = 0..m. f $ (v ! j)) = f $ k * (f $ 0) ^ m" .
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2464
  } note A = this
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2465
  
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2466
  have "(f ^ Suc m) $ k = (\<Sum>v\<in>natpermute k (m + 1). \<Prod>j = 0..m. f $ v ! j)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2467
    by (rule fps_power_nth_Suc)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2468
  also have "natpermute k (m+1) = A \<union> B" unfolding A_def B_def by blast
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2469
  also have "(\<Sum>v\<in>\<dots>. \<Prod>j = 0..m. f $ (v ! j)) = 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2470
               (\<Sum>v\<in>A. \<Prod>j = 0..m. f $ (v ! j)) + (\<Sum>v\<in>B. \<Prod>j = 0..m. f $ (v ! j))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2471
    by (intro sum.union_disjoint) simp_all   
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2472
  also have "(\<Sum>v\<in>A. \<Prod>j = 0..m. f $ (v ! j)) = of_nat (Suc m) * (f $ k * (f $ 0) ^ m)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2473
    by (simp add: A card_A)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2474
  finally show ?thesis by (simp add: B_def)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2475
qed 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2476
  
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2477
lemma fps_power_Suc_eqD:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2478
  fixes f g :: "'a :: {idom,semiring_char_0} fps"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2479
  assumes "f ^ Suc m = g ^ Suc m" "f $ 0 = g $ 0" "f $ 0 \<noteq> 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2480
  shows   "f = g"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2481
proof (rule fps_ext)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2482
  fix k :: nat
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2483
  show "f $ k = g $ k"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2484
  proof (induction k rule: less_induct)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2485
    case (less k)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2486
    show ?case
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2487
    proof (cases "k = 0")
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2488
      case False
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2489
      let ?h = "\<lambda>f. (\<Sum>v | v \<in> natpermute k (m + 1) \<and> k \<notin> set v. \<Prod>j = 0..m. f $ v ! j)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2490
      from False fps_power_Suc_nth[of k f m] fps_power_Suc_nth[of k g m]
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2491
        have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h f =
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2492
                g $ k * (of_nat (Suc m) * (f $ 0) ^ m) + ?h g" using assms 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2493
        by (simp add: mult_ac del: power_Suc of_nat_Suc)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2494
      also have "v ! i < k" if "v \<in> {v\<in>natpermute k (m+1). k \<notin> set v}" "i \<le> m" for v i
66311
037aaa0b6daf added lemmas
nipkow
parents: 66089
diff changeset
  2495
        using that elem_le_sum_list[of i v] unfolding natpermute_def
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2496
        by (auto simp: set_conv_nth dest!: spec[of _ i])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2497
      hence "?h f = ?h g"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2498
        by (intro sum.cong refl prod.cong less lessI) (auto simp: natpermute_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2499
      finally have "f $ k * (of_nat (Suc m) * (f $ 0) ^ m) = g $ k * (of_nat (Suc m) * (f $ 0) ^ m)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2500
        by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2501
      with assms show "f $ k = g $ k" 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2502
        by (subst (asm) mult_right_cancel) (auto simp del: of_nat_Suc)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2503
    qed (simp_all add: assms)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2504
  qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2505
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2506
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2507
lemma fps_power_Suc_eqD':
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2508
  fixes f g :: "'a :: {idom,semiring_char_0} fps"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2509
  assumes "f ^ Suc m = g ^ Suc m" "f $ subdegree f = g $ subdegree g"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2510
  shows   "f = g"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2511
proof (cases "f = 0")
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2512
  case False
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2513
  have "Suc m * subdegree f = subdegree (f ^ Suc m)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2514
    by (rule subdegree_power [symmetric])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2515
  also have "f ^ Suc m = g ^ Suc m" by fact
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2516
  also have "subdegree \<dots> = Suc m * subdegree g" by (rule subdegree_power)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2517
  finally have [simp]: "subdegree f = subdegree g"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2518
    by (subst (asm) Suc_mult_cancel1)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2519
  have "fps_shift (subdegree f) f * fps_X ^ subdegree f = f"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2520
    by (rule subdegree_decompose [symmetric])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2521
  also have "\<dots> ^ Suc m = g ^ Suc m" by fact
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  2522
  also have "g = fps_shift (subdegree g) g * fps_X ^ subdegree g"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2523
    by (rule subdegree_decompose)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2524
  also have "subdegree f = subdegree g" by fact
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2525
  finally have "fps_shift (subdegree g) f ^ Suc m = fps_shift (subdegree g) g ^ Suc m"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2526
    by (simp add: algebra_simps power_mult_distrib del: power_Suc)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2527
  hence "fps_shift (subdegree g) f = fps_shift (subdegree g) g"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2528
    by (rule fps_power_Suc_eqD) (insert assms False, auto)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2529
  with subdegree_decompose[of f] subdegree_decompose[of g] show ?thesis by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2530
qed (insert assms, simp_all)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2531
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2532
lemma fps_power_eqD':
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2533
  fixes f g :: "'a :: {idom,semiring_char_0} fps"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2534
  assumes "f ^ m = g ^ m" "f $ subdegree f = g $ subdegree g" "m > 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2535
  shows   "f = g"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2536
  using fps_power_Suc_eqD'[of f "m-1" g] assms by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2537
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2538
lemma fps_power_eqD:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2539
  fixes f g :: "'a :: {idom,semiring_char_0} fps"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2540
  assumes "f ^ m = g ^ m" "f $ 0 = g $ 0" "f $ 0 \<noteq> 0" "m > 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2541
  shows   "f = g"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2542
  by (rule fps_power_eqD'[of f m g]) (insert assms, simp_all)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  2543
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2544
lemma fps_compose_inj_right:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2545
  assumes a0: "a$0 = (0::'a::idom)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2546
    and a1: "a$1 \<noteq> 0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2547
  shows "(b oo a = c oo a) \<longleftrightarrow> b = c"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2548
  (is "?lhs \<longleftrightarrow>?rhs")
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2549
proof
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2550
  show ?lhs if ?rhs using that by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2551
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2552
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2553
    have "b$n = c$n" for n
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2554
    proof (induct n rule: nat_less_induct)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2555
      fix n
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2556
      assume H: "\<forall>m<n. b$m = c$m"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2557
      show "b$n = c$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2558
      proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2559
        case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2560
        from \<open>?lhs\<close> have "(b oo a)$n = (c oo a)$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2561
          by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2562
        then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2563
          using 0 by (simp add: fps_compose_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2564
      next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2565
        case (Suc n1)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2566
        have f: "finite {0 .. n1}" "finite {n}" by simp_all
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2567
        have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using Suc by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2568
        have d: "{0 .. n1} \<inter> {n} = {}" using Suc by auto
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2569
        have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2570
          apply (rule sum.cong)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2571
          using H Suc
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2572
          apply auto
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2573
          done
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2574
        have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2575
          unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq] seq
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2576
          using startsby_zero_power_nth_same[OF a0]
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2577
          by simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2578
        have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2579
          unfolding fps_compose_nth sum.union_disjoint[OF f d, unfolded eq]
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2580
          using startsby_zero_power_nth_same[OF a0]
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2581
          by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2582
        from \<open>?lhs\<close>[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2583
        show ?thesis by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2584
      qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2585
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2586
    then show ?rhs by (simp add: fps_eq_iff)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2587
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2588
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2589
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2590
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2591
subsection \<open>Radicals\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2592
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2593
declare prod.cong [fundef_cong]
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2594
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2595
function radical :: "(nat \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a::field fps \<Rightarrow> nat \<Rightarrow> 'a"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2596
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2597
  "radical r 0 a 0 = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2598
| "radical r 0 a (Suc n) = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2599
| "radical r (Suc k) a 0 = r (Suc k) (a$0)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2600
| "radical r (Suc k) a (Suc n) =
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2601
    (a$ Suc n - sum (\<lambda>xs. prod (\<lambda>j. radical r (Suc k) a (xs ! j)) {0..k})
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2602
      {xs. xs \<in> natpermute (Suc n) (Suc k) \<and> Suc n \<notin> set xs}) /
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2603
    (of_nat (Suc k) * (radical r (Suc k) a 0)^k)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2604
  by pat_completeness auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2605
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2606
termination radical
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2607
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2608
  let ?R = "measure (\<lambda>(r, k, a, n). n)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2609
  {
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2610
    show "wf ?R" by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2611
  next
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2612
    fix r k a n xs i
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2613
    assume xs: "xs \<in> {xs \<in> natpermute (Suc n) (Suc k). Suc n \<notin> set xs}" and i: "i \<in> {0..k}"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2614
    have False if c: "Suc n \<le> xs ! i"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2615
    proof -
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2616
      from xs i have "xs !i \<noteq> Suc n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2617
        by (auto simp add: in_set_conv_nth natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2618
      with c have c': "Suc n < xs!i" by arith
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2619
      have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2620
        by simp_all
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2621
      have d: "{0 ..< i} \<inter> ({i} \<union> {i+1 ..< Suc k}) = {}" "{i} \<inter> {i+1..< Suc k} = {}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2622
        by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2623
      have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2624
        using i by auto
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2625
      from xs have "Suc n = sum_list xs"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2626
        by (simp add: natpermute_def)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2627
      also have "\<dots> = sum (nth xs) {0..<Suc k}" using xs
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2628
        by (simp add: natpermute_def sum_list_sum_nth)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2629
      also have "\<dots> = xs!i + sum (nth xs) {0..<i} + sum (nth xs) {i+1..<Suc k}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2630
        unfolding eqs  sum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2631
        unfolding sum.union_disjoint[OF fths(2) fths(3) d(2)]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2632
        by simp
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2633
      finally show ?thesis using c' by simp
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2634
    qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2635
    then show "((r, Suc k, a, xs!i), r, Suc k, a, Suc n) \<in> ?R"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2636
      apply auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2637
      apply (metis not_less)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2638
      done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2639
  next
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2640
    fix r k a n
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2641
    show "((r, Suc k, a, 0), r, Suc k, a, Suc n) \<in> ?R" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2642
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2643
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2644
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2645
definition "fps_radical r n a = Abs_fps (radical r n a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2646
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2647
lemma fps_radical0[simp]: "fps_radical r 0 a = 1"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2648
  apply (auto simp add: fps_eq_iff fps_radical_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2649
  apply (case_tac n)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2650
  apply auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2651
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2652
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2653
lemma fps_radical_nth_0[simp]: "fps_radical r n a $ 0 = (if n = 0 then 1 else r n (a$0))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2654
  by (cases n) (simp_all add: fps_radical_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2655
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2656
lemma fps_radical_power_nth[simp]:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2657
  assumes r: "(r k (a$0)) ^ k = a$0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2658
  shows "fps_radical r k a ^ k $ 0 = (if k = 0 then 1 else a$0)"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2659
proof (cases k)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2660
  case 0
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2661
  then show ?thesis by simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2662
next
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2663
  case (Suc h)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2664
  have eq1: "fps_radical r k a ^ k $ 0 = (\<Prod>j\<in>{0..h}. fps_radical r k a $ (replicate k 0) ! j)"
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2665
    unfolding fps_power_nth Suc by simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2666
  also have "\<dots> = (\<Prod>j\<in>{0..h}. r k (a$0))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2667
    apply (rule prod.cong)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2668
    apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2669
    using Suc
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2670
    apply (subgoal_tac "replicate k 0 ! x = 0")
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2671
    apply (auto intro: nth_replicate simp del: replicate.simps)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2672
    done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2673
  also have "\<dots> = a$0"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2674
    using r Suc by (simp add: prod_constant)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2675
  finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2676
    using Suc by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2677
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2678
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2679
lemma power_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2680
  fixes a:: "'a::field_char_0 fps"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2681
  assumes a0: "a$0 \<noteq> 0"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2682
  shows "(r (Suc k) (a$0)) ^ Suc k = a$0 \<longleftrightarrow> (fps_radical r (Suc k) a) ^ (Suc k) = a"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2683
    (is "?lhs \<longleftrightarrow> ?rhs")
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2684
proof
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2685
  let ?r = "fps_radical r (Suc k) a"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2686
  show ?rhs if r0: ?lhs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2687
  proof -
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2688
    from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2689
    have "?r ^ Suc k $ z = a$z" for z
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2690
    proof (induct z rule: nat_less_induct)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2691
      fix n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2692
      assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2693
      show "?r ^ Suc k $ n = a $n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2694
      proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2695
        case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2696
        then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2697
          using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2698
      next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2699
        case (Suc n1)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2700
        then have "n \<noteq> 0" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2701
        let ?Pnk = "natpermute n (k + 1)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2702
        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2703
        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2704
        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2705
        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2706
        have f: "finite ?Pnkn" "finite ?Pnknn"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2707
          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2708
          by (metis natpermute_finite)+
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2709
        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2710
        have "sum ?f ?Pnkn = sum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2711
        proof (rule sum.cong)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2712
          fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2713
          let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2714
            fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2715
          from v obtain i where i: "i \<in> {0..k}" "v = (replicate (k+1) 0) [i:= n]"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2716
            unfolding natpermute_contain_maximal by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2717
          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2718
              (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2719
            apply (rule prod.cong, simp)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2720
            using i r0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2721
            apply (simp del: replicate.simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2722
            done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2723
          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2724
            using i r0 by (simp add: prod_gen_delta)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2725
          finally show ?ths .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2726
        qed rule
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2727
        then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2728
          by (simp add: natpermute_max_card[OF \<open>n \<noteq> 0\<close>, simplified])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2729
        also have "\<dots> = a$n - sum ?f ?Pnknn"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2730
          unfolding Suc using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2731
        finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" .
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2732
        have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2733
          unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] ..
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2734
        also have "\<dots> = a$n" unfolding fn by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2735
        finally show ?thesis .
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2736
      qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2737
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2738
    then show ?thesis using r0 by (simp add: fps_eq_iff)
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2739
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2740
  show ?lhs if ?rhs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2741
  proof -
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2742
    from that have "((fps_radical r (Suc k) a) ^ (Suc k))$0 = a$0"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2743
      by simp
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2744
    then show ?thesis
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2745
      unfolding fps_power_nth_Suc
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2746
      by (simp add: prod_constant del: replicate.simps)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2747
  qed
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2748
qed
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2749
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2750
(*
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2751
lemma power_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2752
  fixes a:: "'a::field_char_0 fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2753
  assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0" and a0: "a$0 \<noteq> 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2754
  shows "(fps_radical r (Suc k) a) ^ (Suc k) = a"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2755
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2756
  let ?r = "fps_radical r (Suc k) a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2757
  from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2758
  {fix z have "?r ^ Suc k $ z = a$z"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2759
    proof(induct z rule: nat_less_induct)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2760
      fix n assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2761
      {assume "n = 0" then have "?r ^ Suc k $ n = a $n"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2762
          using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2763
      moreover
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2764
      {fix n1 assume n1: "n = Suc n1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2765
        have fK: "finite {0..k}" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2766
        have nz: "n \<noteq> 0" using n1 by arith
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2767
        let ?Pnk = "natpermute n (k + 1)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2768
        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2769
        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2770
        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2771
        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2772
        have f: "finite ?Pnkn" "finite ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2773
          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2774
          by (metis natpermute_finite)+
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2775
        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2776
        have "sum ?f ?Pnkn = sum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2777
        proof(rule sum.cong2)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2778
          fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2779
          let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2780
          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2781
            unfolding natpermute_contain_maximal by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2782
          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2783
            apply (rule prod.cong, simp)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2784
            using i r0 by (simp del: replicate.simps)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2785
          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2786
            unfolding prod_gen_delta[OF fK] using i r0 by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2787
          finally show ?ths .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2788
        qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2789
        then have "sum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2790
          by (simp add: natpermute_max_card[OF nz, simplified])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2791
        also have "\<dots> = a$n - sum ?f ?Pnknn"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2792
          unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2793
        finally have fn: "sum ?f ?Pnkn = a$n - sum ?f ?Pnknn" .
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2794
        have "(?r ^ Suc k)$n = sum ?f ?Pnkn + sum ?f ?Pnknn"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2795
          unfolding fps_power_nth_Suc sum.union_disjoint[OF f d, unfolded eq] ..
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2796
        also have "\<dots> = a$n" unfolding fn by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2797
        finally have "?r ^ Suc k $ n = a $n" .}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2798
      ultimately  show "?r ^ Suc k $ n = a $n" by (cases n, auto)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2799
  qed }
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2800
  then show ?thesis by (simp add: fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2801
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2802
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2803
*)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2804
lemma eq_divide_imp':
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2805
  fixes c :: "'a::field"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2806
  shows "c \<noteq> 0 \<Longrightarrow> a * c = b \<Longrightarrow> a = b / c"
56480
093ea91498e6 field_simps: better support for negation and division, and power
hoelzl
parents: 56479
diff changeset
  2807
  by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2808
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2809
lemma radical_unique:
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2810
  assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2811
    and a0: "r (Suc k) (b$0 ::'a::field_char_0) = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2812
    and b0: "b$0 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2813
  shows "a^(Suc k) = b \<longleftrightarrow> a = fps_radical r (Suc k) b"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2814
    (is "?lhs \<longleftrightarrow> ?rhs" is "_ \<longleftrightarrow> a = ?r")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2815
proof
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2816
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2817
    using that using power_radical[OF b0, of r k, unfolded r0] by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2818
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2819
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2820
    have r00: "r (Suc k) (b$0) \<noteq> 0" using b0 r0 by auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2821
    have ceq: "card {0..k} = Suc k" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2822
    from a0 have a0r0: "a$0 = ?r$0" by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2823
    have "a $ n = ?r $ n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2824
    proof (induct n rule: nat_less_induct)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2825
      fix n
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2826
      assume h: "\<forall>m<n. a$m = ?r $m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2827
      show "a$n = ?r $ n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2828
      proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2829
        case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2830
        then show ?thesis using a0 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2831
      next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2832
        case (Suc n1)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2833
        have fK: "finite {0..k}" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2834
        have nz: "n \<noteq> 0" using Suc by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2835
        let ?Pnk = "natpermute n (Suc k)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2836
        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2837
        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2838
        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2839
        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2840
        have f: "finite ?Pnkn" "finite ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2841
          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2842
          by (metis natpermute_finite)+
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2843
        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2844
        let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2845
        have "sum ?g ?Pnkn = sum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2846
        proof (rule sum.cong)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2847
          fix v
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2848
          assume v: "v \<in> {xs \<in> natpermute n (Suc k). n \<in> set xs}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2849
          let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
69085
9999d7823b8f updated to new list_update precedence
nipkow
parents: 69064
diff changeset
  2850
          from v obtain i where i: "i \<in> {0..k}" "v = (replicate (k+1) 0) [i:= n]"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2851
            unfolding Suc_eq_plus1 natpermute_contain_maximal
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2852
            by (auto simp del: replicate.simps)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2853
          have "(\<Prod>j\<in>{0..k}. a $ v ! j) = (\<Prod>j\<in>{0..k}. if j = i then a $ n else r (Suc k) (b$0))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2854
            apply (rule prod.cong, simp)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2855
            using i a0
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2856
            apply (simp del: replicate.simps)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2857
            done
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2858
          also have "\<dots> = a $ n * (?r $ 0)^k"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2859
            using i by (simp add: prod_gen_delta)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2860
          finally show ?ths .
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2861
        qed rule
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2862
        then have th0: "sum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2863
          by (simp add: natpermute_max_card[OF nz, simplified])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2864
        have th1: "sum ?g ?Pnknn = sum ?f ?Pnknn"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2865
        proof (rule sum.cong, rule refl, rule prod.cong, simp)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2866
          fix xs i
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2867
          assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2868
          have False if c: "n \<le> xs ! i"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2869
          proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2870
            from xs i have "xs ! i \<noteq> n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2871
              by (auto simp add: in_set_conv_nth natpermute_def)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2872
            with c have c': "n < xs!i" by arith
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2873
            have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2874
              by simp_all
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2875
            have d: "{0 ..< i} \<inter> ({i} \<union> {i+1 ..< Suc k}) = {}" "{i} \<inter> {i+1..< Suc k} = {}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2876
              by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2877
            have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2878
              using i by auto
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63589
diff changeset
  2879
            from xs have "n = sum_list xs"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2880
              by (simp add: natpermute_def)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2881
            also have "\<dots> = sum (nth xs) {0..<Suc k}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2882
              using xs by (simp add: natpermute_def sum_list_sum_nth)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2883
            also have "\<dots> = xs!i + sum (nth xs) {0..<i} + sum (nth xs) {i+1..<Suc k}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2884
              unfolding eqs  sum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2885
              unfolding sum.union_disjoint[OF fths(2) fths(3) d(2)]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2886
              by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2887
            finally show ?thesis using c' by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2888
          qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2889
          then have thn: "xs!i < n" by presburger
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2890
          from h[rule_format, OF thn] show "a$(xs !i) = ?r$(xs!i)" .
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2891
        qed
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2892
        have th00: "\<And>x::'a. of_nat (Suc k) * (x * inverse (of_nat (Suc k))) = x"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  2893
          by (simp add: field_simps del: of_nat_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2894
        from \<open>?lhs\<close> have "b$n = a^Suc k $ n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2895
          by (simp add: fps_eq_iff)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2896
        also have "a ^ Suc k$n = sum ?g ?Pnkn + sum ?g ?Pnknn"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2897
          unfolding fps_power_nth_Suc
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2898
          using sum.union_disjoint[OF f d, unfolded Suc_eq_plus1[symmetric],
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2899
            unfolded eq, of ?g] by simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2900
        also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + sum ?f ?Pnknn"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2901
          unfolding th0 th1 ..
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2902
        finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - sum ?f ?Pnknn"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2903
          by simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  2904
        then have "a$n = (b$n - sum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2905
          apply -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2906
          apply (rule eq_divide_imp')
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2907
          using r00
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2908
          apply (simp del: of_nat_Suc)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  2909
          apply (simp add: ac_simps)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2910
          done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2911
        then show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2912
          apply (simp del: of_nat_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2913
          unfolding fps_radical_def Suc
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2914
          apply (simp add: field_simps Suc th00 del: of_nat_Suc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2915
          done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2916
      qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2917
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2918
    then show ?rhs by (simp add: fps_eq_iff)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2919
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2920
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2921
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2922
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2923
lemma radical_power:
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2924
  assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2925
    and a0: "(a$0 :: 'a::field_char_0) \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2926
  shows "(fps_radical r (Suc k) (a ^ Suc k)) = a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2927
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2928
  let ?ak = "a^ Suc k"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2929
  have ak0: "?ak $ 0 = (a$0) ^ Suc k"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2930
    by (simp add: fps_nth_power_0 del: power_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2931
  from r0 have th0: "r (Suc k) (a ^ Suc k $ 0) ^ Suc k = a ^ Suc k $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2932
    using ak0 by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2933
  from r0 ak0 have th1: "r (Suc k) (a ^ Suc k $ 0) = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2934
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2935
  from ak0 a0 have ak00: "?ak $ 0 \<noteq>0 "
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2936
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2937
  from radical_unique[of r k ?ak a, OF th0 th1 ak00] show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2938
    by metis
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2939
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2940
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2941
lemma fps_deriv_radical:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2942
  fixes a :: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2943
  assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2944
    and a0: "a$0 \<noteq> 0"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2945
  shows "fps_deriv (fps_radical r (Suc k) a) =
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2946
    fps_deriv a / (fps_const (of_nat (Suc k)) * (fps_radical r (Suc k) a) ^ k)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2947
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2948
  let ?r = "fps_radical r (Suc k) a"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2949
  let ?w = "(fps_const (of_nat (Suc k)) * ?r ^ k)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2950
  from a0 r0 have r0': "r (Suc k) (a$0) \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2951
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2952
  from r0' have w0: "?w $ 0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2953
    by (simp del: of_nat_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2954
  note th0 = inverse_mult_eq_1[OF w0]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2955
  let ?iw = "inverse ?w"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2956
  from iffD1[OF power_radical[of a r], OF a0 r0]
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2957
  have "fps_deriv (?r ^ Suc k) = fps_deriv a"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2958
    by simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2959
  then have "fps_deriv ?r * ?w = fps_deriv a"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  2960
    by (simp add: fps_deriv_power ac_simps del: power_Suc)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2961
  then have "?iw * fps_deriv ?r * ?w = ?iw * fps_deriv a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2962
    by simp
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  2963
  with a0 r0 have "fps_deriv ?r * (?iw * ?w) = fps_deriv a / ?w"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  2964
    by (subst fps_divide_unit) (auto simp del: of_nat_Suc)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2965
  then show ?thesis unfolding th0 by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2966
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2967
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2968
lemma radical_mult_distrib:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2969
  fixes a :: "'a::field_char_0 fps"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2970
  assumes k: "k > 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2971
    and ra0: "r k (a $ 0) ^ k = a $ 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2972
    and rb0: "r k (b $ 0) ^ k = b $ 0"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2973
    and a0: "a $ 0 \<noteq> 0"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2974
    and b0: "b $ 0 \<noteq> 0"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2975
  shows "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0) \<longleftrightarrow>
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2976
    fps_radical r k (a * b) = fps_radical r k a * fps_radical r k b"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2977
    (is "?lhs \<longleftrightarrow> ?rhs")
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2978
proof
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2979
  show ?rhs if r0': ?lhs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2980
  proof -
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2981
    from r0' have r0: "(r k ((a * b) $ 0)) ^ k = (a * b) $ 0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2982
      by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2983
    show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2984
    proof (cases k)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2985
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2986
      then show ?thesis using r0' by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2987
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2988
      case (Suc h)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2989
      let ?ra = "fps_radical r (Suc h) a"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2990
      let ?rb = "fps_radical r (Suc h) b"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2991
      have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2992
        using r0' Suc by (simp add: fps_mult_nth)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2993
      have ab0: "(a*b) $ 0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2994
        using a0 b0 by (simp add: fps_mult_nth)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2995
      from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded Suc] th0 ab0, symmetric]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2996
        iffD1[OF power_radical[of _ r], OF a0 ra0[unfolded Suc]] iffD1[OF power_radical[of _ r], OF b0 rb0[unfolded Suc]] Suc r0'
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2997
      show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2998
        by (auto simp add: power_mult_distrib simp del: power_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2999
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3000
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3001
  show ?lhs if ?rhs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3002
  proof -
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3003
    from that have "(fps_radical r k (a * b)) $ 0 = (fps_radical r k a * fps_radical r k b) $ 0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3004
      by simp
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3005
    then show ?thesis
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3006
      using k by (simp add: fps_mult_nth)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3007
  qed
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3008
qed
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3009
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3010
(*
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3011
lemma radical_mult_distrib:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3012
  fixes a:: "'a::field_char_0 fps"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3013
  assumes
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3014
  ra0: "r k (a $ 0) ^ k = a $ 0"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3015
  and rb0: "r k (b $ 0) ^ k = b $ 0"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3016
  and r0': "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3017
  and a0: "a$0 \<noteq> 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3018
  and b0: "b$0 \<noteq> 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3019
  shows "fps_radical r (k) (a*b) = fps_radical r (k) a * fps_radical r (k) (b)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3020
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3021
  from r0' have r0: "(r (k) ((a*b)$0)) ^ k = (a*b)$0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3022
    by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3023
  {assume "k=0" then have ?thesis by simp}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3024
  moreover
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3025
  {fix h assume k: "k = Suc h"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3026
  let ?ra = "fps_radical r (Suc h) a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3027
  let ?rb = "fps_radical r (Suc h) b"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3028
  have th0: "r (Suc h) ((a * b) $ 0) = (fps_radical r (Suc h) a * fps_radical r (Suc h) b) $ 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3029
    using r0' k by (simp add: fps_mult_nth)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3030
  have ab0: "(a*b) $ 0 \<noteq> 0" using a0 b0 by (simp add: fps_mult_nth)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3031
  from radical_unique[of r h "a*b" "fps_radical r (Suc h) a * fps_radical r (Suc h) b", OF r0[unfolded k] th0 ab0, symmetric]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3032
    power_radical[of r, OF ra0[unfolded k] a0] power_radical[of r, OF rb0[unfolded k] b0] k
30273
ecd6f0ca62ea declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
huffman
parents: 29915
diff changeset
  3033
  have ?thesis by (auto simp add: power_mult_distrib simp del: power_Suc)}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3034
ultimately show ?thesis by (cases k, auto)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3035
qed
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3036
*)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3037
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3038
lemma fps_divide_1 [simp]: "(a :: 'a::field fps) / 1 = a"
64240
eabf80376aab more standardized names
haftmann
parents: 63918
diff changeset
  3039
  by (fact div_by_1)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3040
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3041
lemma radical_divide:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3042
  fixes a :: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3043
  assumes kp: "k > 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3044
    and ra0: "(r k (a $ 0)) ^ k = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3045
    and rb0: "(r k (b $ 0)) ^ k = b $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3046
    and a0: "a$0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3047
    and b0: "b$0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3048
  shows "r k ((a $ 0) / (b$0)) = r k (a$0) / r k (b $ 0) \<longleftrightarrow>
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3049
    fps_radical r k (a/b) = fps_radical r k a / fps_radical r k b"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3050
  (is "?lhs = ?rhs")
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3051
proof
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3052
  let ?r = "fps_radical r k"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3053
  from kp obtain h where k: "k = Suc h"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3054
    by (cases k) auto
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3055
  have ra0': "r k (a$0) \<noteq> 0" using a0 ra0 k by auto
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3056
  have rb0': "r k (b$0) \<noteq> 0" using b0 rb0 k by auto
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3057
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3058
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3059
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3060
    from that have "?r (a/b) $ 0 = (?r a / ?r b)$0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3061
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3062
    then show ?thesis
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3063
      using k a0 b0 rb0' by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3064
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3065
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3066
  proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3067
    from a0 b0 have ab0[simp]: "(a/b)$0 = a$0 / b$0"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3068
      by (simp add: fps_divide_def fps_mult_nth divide_inverse fps_inverse_def)
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3069
    have th0: "r k ((a/b)$0) ^ k = (a/b)$0"
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
  3070
      by (simp add: \<open>?lhs\<close> power_divide ra0 rb0)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3071
    from a0 b0 ra0' rb0' kp \<open>?lhs\<close>
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3072
    have th1: "r k ((a / b) $ 0) = (fps_radical r k a / fps_radical r k b) $ 0"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3073
      by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def divide_inverse)
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3074
    from a0 b0 ra0' rb0' kp have ab0': "(a / b) $ 0 \<noteq> 0"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3075
      by (simp add: fps_divide_unit fps_mult_nth fps_inverse_def nonzero_imp_inverse_nonzero)
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3076
    note tha[simp] = iffD1[OF power_radical[where r=r and k=h], OF a0 ra0[unfolded k], unfolded k[symmetric]]
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3077
    note thb[simp] = iffD1[OF power_radical[where r=r and k=h], OF b0 rb0[unfolded k], unfolded k[symmetric]]
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3078
    from b0 rb0' have th2: "(?r a / ?r b)^k = a/b"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3079
      by (simp add: fps_divide_unit power_mult_distrib fps_inverse_power[symmetric])
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  3080
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3081
    from iffD1[OF radical_unique[where r=r and a="?r a / ?r b" and b="a/b" and k=h], symmetric, unfolded k[symmetric], OF th0 th1 ab0' th2]
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3082
    show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3083
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3084
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3085
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3086
lemma radical_inverse:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3087
  fixes a :: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3088
  assumes k: "k > 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3089
    and ra0: "r k (a $ 0) ^ k = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3090
    and r1: "(r k 1)^k = 1"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3091
    and a0: "a$0 \<noteq> 0"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3092
  shows "r k (inverse (a $ 0)) = r k 1 / (r k (a $ 0)) \<longleftrightarrow>
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3093
    fps_radical r k (inverse a) = fps_radical r k 1 / fps_radical r k a"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3094
  using radical_divide[where k=k and r=r and a=1 and b=a, OF k ] ra0 r1 a0
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3095
  by (simp add: divide_inverse fps_divide_def)
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  3096
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3097
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3098
subsection \<open>Derivative of composition\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3099
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3100
lemma fps_compose_deriv:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3101
  fixes a :: "'a::idom fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3102
  assumes b0: "b$0 = 0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3103
  shows "fps_deriv (a oo b) = ((fps_deriv a) oo b) * fps_deriv b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3104
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3105
  have "(fps_deriv (a oo b))$n = (((fps_deriv a) oo b) * (fps_deriv b)) $n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3106
  proof -
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3107
    have "(fps_deriv (a oo b))$n = sum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3108
      by (simp add: fps_compose_def field_simps sum_distrib_left del: of_nat_Suc)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3109
    also have "\<dots> = sum (\<lambda>i. a$i * ((fps_const (of_nat i)) * (fps_deriv b * (b^(i - 1))))$n) {0.. Suc n}"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3110
      by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3111
    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3112
      unfolding fps_mult_left_const_nth  by (simp add: field_simps)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3113
    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (sum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3114
      unfolding fps_mult_nth ..
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3115
    also have "\<dots> = sum (\<lambda>i. of_nat i * a$i * (sum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3116
      apply (rule sum.mono_neutral_right)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3117
      apply (auto simp add: mult_delta_left sum.delta not_le)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3118
      done
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3119
    also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3120
      unfolding fps_deriv_nth
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3121
      by (rule sum.reindex_cong [of Suc]) (auto simp add: mult.assoc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3122
    finally have th0: "(fps_deriv (a oo b))$n =
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3123
      sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3124
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3125
    have "(((fps_deriv a) oo b) * (fps_deriv b))$n = sum (\<lambda>i. (fps_deriv b)$ (n - i) * ((fps_deriv a) oo b)$i) {0..n}"
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  3126
      unfolding fps_mult_nth by (simp add: ac_simps)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3127
    also have "\<dots> = sum (\<lambda>i. sum (\<lambda>j. of_nat (n - i +1) * b$(n - i + 1) * of_nat (j + 1) * a$(j+1) * (b^j)$i) {0..n}) {0..n}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3128
      unfolding fps_deriv_nth fps_compose_nth sum_distrib_left mult.assoc
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3129
      apply (rule sum.cong)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3130
      apply (rule refl)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3131
      apply (rule sum.mono_neutral_left)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3132
      apply (simp_all add: subset_eq)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3133
      apply clarify
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3134
      apply (subgoal_tac "b^i$x = 0")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3135
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3136
      apply (rule startsby_zero_power_prefix[OF b0, rule_format])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3137
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3138
      done
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3139
    also have "\<dots> = sum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (sum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3140
      unfolding sum_distrib_left
66804
3f9bb52082c4 avoid name clashes on interpretation of abstract locales
haftmann
parents: 66550
diff changeset
  3141
      apply (subst sum.swap)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3142
      apply (rule sum.cong, rule refl)+
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3143
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3144
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3145
    finally show ?thesis
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3146
      unfolding th0 by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3147
  qed
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3148
  then show ?thesis by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3149
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3150
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3151
lemma fps_mult_fps_X_plus_1_nth:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3152
  "((1+fps_X)*a) $n = (if n = 0 then (a$n :: 'a::comm_ring_1) else a$n + a$(n - 1))"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3153
proof (cases n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3154
  case 0
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3155
  then show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3156
    by (simp add: fps_mult_nth)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3157
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3158
  case (Suc m)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3159
  have "((1 + fps_X)*a) $ n = sum (\<lambda>i. (1 + fps_X) $ i * a $ (n - i)) {0..n}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3160
    by (simp add: fps_mult_nth)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3161
  also have "\<dots> = sum (\<lambda>i. (1+fps_X)$i * a$(n-i)) {0.. 1}"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3162
    unfolding Suc by (rule sum.mono_neutral_right) auto
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3163
  also have "\<dots> = (if n = 0 then (a$n :: 'a::comm_ring_1) else a$n + a$(n - 1))"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3164
    by (simp add: Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3165
  finally show ?thesis .
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3166
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3167
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3168
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3169
subsection \<open>Finite FPS (i.e. polynomials) and fps_X\<close>
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3170
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3171
lemma fps_poly_sum_fps_X:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3172
  assumes "\<forall>i > n. a$i = (0::'a::comm_ring_1)"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3173
  shows "a = sum (\<lambda>i. fps_const (a$i) * fps_X^i) {0..n}" (is "a = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3174
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3175
  have "a$i = ?r$i" for i
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3176
    unfolding fps_sum_nth fps_mult_left_const_nth fps_X_power_nth
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3177
    by (simp add: mult_delta_right sum.delta' assms)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3178
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3179
    unfolding fps_eq_iff by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3180
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3181
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3182
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3183
subsection \<open>Compositional inverses\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3184
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3185
fun compinv :: "'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3186
where
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3187
  "compinv a 0 = fps_X$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3188
| "compinv a (Suc n) =
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3189
    (fps_X$ Suc n - sum (\<lambda>i. (compinv a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3190
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3191
definition "fps_inv a = Abs_fps (compinv a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3192
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3193
lemma fps_inv:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3194
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3195
    and a1: "a$1 \<noteq> 0"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3196
  shows "fps_inv a oo a = fps_X"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3197
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3198
  let ?i = "fps_inv a oo a"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3199
  have "?i $n = fps_X$n" for n
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3200
  proof (induct n rule: nat_less_induct)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3201
    fix n
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3202
    assume h: "\<forall>m<n. ?i$m = fps_X$m"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3203
    show "?i $ n = fps_X$n"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3204
    proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3205
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3206
      then show ?thesis using a0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3207
        by (simp add: fps_compose_nth fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3208
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3209
      case (Suc n1)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3210
      have "?i $ n = sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3211
        by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3212
      also have "\<dots> = sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} +
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3213
        (fps_X$ Suc n1 - sum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3214
        using a0 a1 Suc by (simp add: fps_inv_def)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3215
      also have "\<dots> = fps_X$n" using Suc by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3216
      finally show ?thesis .
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3217
    qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3218
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3219
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3220
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3221
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3222
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3223
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3224
fun gcompinv :: "'a fps \<Rightarrow> 'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3225
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3226
  "gcompinv b a 0 = b$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3227
| "gcompinv b a (Suc n) =
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3228
    (b$ Suc n - sum (\<lambda>i. (gcompinv b a i) * (a^i)$Suc n) {0 .. n}) / (a$1) ^ Suc n"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3229
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3230
definition "fps_ginv b a = Abs_fps (gcompinv b a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3231
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3232
lemma fps_ginv:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3233
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3234
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3235
  shows "fps_ginv b a oo a = b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3236
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3237
  let ?i = "fps_ginv b a oo a"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3238
  have "?i $n = b$n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3239
  proof (induct n rule: nat_less_induct)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3240
    fix n
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3241
    assume h: "\<forall>m<n. ?i$m = b$m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3242
    show "?i $ n = b$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3243
    proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3244
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3245
      then show ?thesis using a0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3246
        by (simp add: fps_compose_nth fps_ginv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3247
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3248
      case (Suc n1)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3249
      have "?i $ n = sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3250
        by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3251
      also have "\<dots> = sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} +
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3252
        (b$ Suc n1 - sum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3253
        using a0 a1 Suc by (simp add: fps_ginv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3254
      also have "\<dots> = b$n" using Suc by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3255
      finally show ?thesis .
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3256
    qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3257
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3258
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3259
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3260
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3261
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3262
lemma fps_inv_ginv: "fps_inv = fps_ginv fps_X"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  3263
  apply (auto simp add: fun_eq_iff fps_eq_iff fps_inv_def fps_ginv_def)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3264
  apply (induct_tac n rule: nat_less_induct)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3265
  apply auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3266
  apply (case_tac na)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3267
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3268
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3269
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3270
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3271
lemma fps_compose_1[simp]: "1 oo a = 1"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3272
  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3273
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3274
lemma fps_compose_0[simp]: "0 oo a = 0"
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
  3275
  by (simp add: fps_eq_iff fps_compose_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3276
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
  3277
lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a $ 0)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3278
  by (auto simp add: fps_eq_iff fps_compose_nth power_0_left sum.neutral)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3279
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3280
lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3281
  by (simp add: fps_eq_iff fps_compose_nth field_simps sum.distrib)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3282
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3283
lemma fps_compose_sum_distrib: "(sum f S) oo a = sum (\<lambda>i. f i oo a) S"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3284
proof (cases "finite S")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3285
  case True
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3286
  show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3287
  proof (rule finite_induct[OF True])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3288
    show "sum f {} oo a = (\<Sum>i\<in>{}. f i oo a)"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3289
      by simp
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3290
  next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3291
    fix x F
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3292
    assume fF: "finite F"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3293
      and xF: "x \<notin> F"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3294
      and h: "sum f F oo a = sum (\<lambda>i. f i oo a) F"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3295
    show "sum f (insert x F) oo a  = sum (\<lambda>i. f i oo a) (insert x F)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3296
      using fF xF h by (simp add: fps_compose_add_distrib)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3297
  qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3298
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3299
  case False
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3300
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3301
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3302
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3303
lemma convolution_eq:
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3304
  "sum (\<lambda>i. a (i :: nat) * b (n - i)) {0 .. n} =
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3305
    sum (\<lambda>(i,j). a i * b j) {(i,j). i \<le> n \<and> j \<le> n \<and> i + j = n}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3306
  by (rule sum.reindex_bij_witness[where i=fst and j="\<lambda>i. (i, n - i)"]) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3307
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3308
lemma product_composition_lemma:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3309
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3310
    and d0: "d$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3311
  shows "((a oo c) * (b oo d))$n =
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3312
    sum (\<lambda>(k,m). a$k * b$m * (c^k * d^m) $ n) {(k,m). k + m \<le> n}"  (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3313
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3314
  let ?S = "{(k::nat, m::nat). k + m \<le> n}"
61943
7fba644ed827 discontinued ASCII replacement syntax <*>;
wenzelm
parents: 61804
diff changeset
  3315
  have s: "?S \<subseteq> {0..n} \<times> {0..n}" by (auto simp add: subset_eq)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3316
  have f: "finite {(k::nat, m::nat). k + m \<le> n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3317
    apply (rule finite_subset[OF s])
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3318
    apply auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3319
    done
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3320
  have "?r =  sum (\<lambda>i. sum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3321
    apply (simp add: fps_mult_nth sum_distrib_left)
66804
3f9bb52082c4 avoid name clashes on interpretation of abstract locales
haftmann
parents: 66550
diff changeset
  3322
    apply (subst sum.swap)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3323
    apply (rule sum.cong)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3324
    apply (auto simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3325
    done
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3326
  also have "\<dots> = ?l"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3327
    apply (simp add: fps_mult_nth fps_compose_nth sum_product)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3328
    apply (rule sum.cong)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3329
    apply (rule refl)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3330
    apply (simp add: sum.cartesian_product mult.assoc)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3331
    apply (rule sum.mono_neutral_right[OF f])
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3332
    apply (simp add: subset_eq)
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3333
    apply presburger
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3334
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3335
    apply (rule ccontr)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3336
    apply (clarsimp simp add: not_le)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3337
    apply (case_tac "x < aa")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3338
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3339
    apply (frule_tac startsby_zero_power_prefix[rule_format, OF c0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3340
    apply blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3341
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3342
    apply (frule_tac startsby_zero_power_prefix[rule_format, OF d0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3343
    apply blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3344
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3345
  finally show ?thesis by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3346
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3347
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3348
lemma product_composition_lemma':
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3349
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3350
    and d0: "d$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3351
  shows "((a oo c) * (b oo d))$n =
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3352
    sum (\<lambda>k. sum (\<lambda>m. a$k * b$m * (c^k * d^m) $ n) {0..n}) {0..n}"  (is "?l = ?r")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3353
  unfolding product_composition_lemma[OF c0 d0]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3354
  unfolding sum.cartesian_product
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3355
  apply (rule sum.mono_neutral_left)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3356
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3357
  apply (clarsimp simp add: subset_eq)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3358
  apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3359
  apply (rule ccontr)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3360
  apply (subgoal_tac "(c^aa * d^ba) $ n = 0")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3361
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3362
  unfolding fps_mult_nth
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3363
  apply (rule sum.neutral)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3364
  apply (clarsimp simp add: not_le)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51107
diff changeset
  3365
  apply (case_tac "x < aa")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3366
  apply (rule startsby_zero_power_prefix[OF c0, rule_format])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3367
  apply simp
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51107
diff changeset
  3368
  apply (subgoal_tac "n - x < ba")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3369
  apply (frule_tac k = "ba" in startsby_zero_power_prefix[OF d0, rule_format])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3370
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3371
  apply arith
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3372
  done
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3373
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3374
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3375
lemma sum_pair_less_iff:
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3376
  "sum (\<lambda>((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \<le> n} =
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3377
    sum (\<lambda>s. sum (\<lambda>i. a i * b (s - i) * c s) {0..s}) {0..n}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3378
  (is "?l = ?r")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3379
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3380
  let ?KM = "{(k,m). k + m \<le> n}"
69325
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69260
diff changeset
  3381
  let ?f = "\<lambda>s. \<Union>i\<in>{0..s}. {(i, s - i)}"
4b6ddc5989fc removed legacy input syntax
haftmann
parents: 69260
diff changeset
  3382
  have th0: "?KM = \<Union> (?f ` {0..n})"
62343
24106dc44def prefer abbreviations for compound operators INFIMUM and SUPREMUM
haftmann
parents: 62102
diff changeset
  3383
    by auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3384
  show "?l = ?r "
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3385
    unfolding th0
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3386
    apply (subst sum.UNION_disjoint)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3387
    apply auto
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3388
    apply (subst sum.UNION_disjoint)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3389
    apply auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3390
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3391
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3392
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3393
lemma fps_compose_mult_distrib_lemma:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3394
  assumes c0: "c$0 = (0::'a::idom)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3395
  shows "((a oo c) * (b oo c))$n = sum (\<lambda>s. sum (\<lambda>i. a$i * b$(s - i) * (c^s) $ n) {0..s}) {0..n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3396
  unfolding product_composition_lemma[OF c0 c0] power_add[symmetric]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3397
  unfolding sum_pair_less_iff[where a = "\<lambda>k. a$k" and b="\<lambda>m. b$m" and c="\<lambda>s. (c ^ s)$n" and n = n] ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3398
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3399
lemma fps_compose_mult_distrib:
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54452
diff changeset
  3400
  assumes c0: "c $ 0 = (0::'a::idom)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54452
diff changeset
  3401
  shows "(a * b) oo c = (a oo c) * (b oo c)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54452
diff changeset
  3402
  apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3403
  apply (simp add: fps_compose_nth fps_mult_nth sum_distrib_right)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3404
  done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3405
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3406
lemma fps_compose_prod_distrib:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3407
  assumes c0: "c$0 = (0::'a::idom)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3408
  shows "prod a S oo c = prod (\<lambda>k. a k oo c) S"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3409
  apply (cases "finite S")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3410
  apply simp_all
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3411
  apply (induct S rule: finite_induct)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3412
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3413
  apply (simp add: fps_compose_mult_distrib[OF c0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3414
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3415
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3416
lemma fps_compose_divide:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3417
  assumes [simp]: "g dvd f" "h $ 0 = 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3418
  shows   "fps_compose f h = fps_compose (f / g :: 'a :: field fps) h * fps_compose g h"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3419
proof -
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3420
  have "f = (f / g) * g" by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3421
  also have "fps_compose \<dots> h = fps_compose (f / g) h * fps_compose g h"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3422
    by (subst fps_compose_mult_distrib) simp_all
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3423
  finally show ?thesis .
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3424
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3425
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3426
lemma fps_compose_divide_distrib:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3427
  assumes "g dvd f" "h $ 0 = 0" "fps_compose g h \<noteq> 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3428
  shows   "fps_compose (f / g :: 'a :: field fps) h = fps_compose f h / fps_compose g h"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3429
  using fps_compose_divide[OF assms(1,2)] assms(3) by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3430
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3431
lemma fps_compose_power:
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3432
  assumes c0: "c$0 = (0::'a::idom)"
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3433
  shows "(a oo c)^n = a^n oo c"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3434
proof (cases n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3435
  case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3436
  then show ?thesis by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3437
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3438
  case (Suc m)
67970
8c012a49293a A couple of new results
paulson <lp15@cam.ac.uk>
parents: 67411
diff changeset
  3439
  have "(\<Prod>n = 0..m. a) oo c = (\<Prod>n = 0..m. a oo c)"
8c012a49293a A couple of new results
paulson <lp15@cam.ac.uk>
parents: 67411
diff changeset
  3440
    using c0 fps_compose_prod_distrib by blast
8c012a49293a A couple of new results
paulson <lp15@cam.ac.uk>
parents: 67411
diff changeset
  3441
  moreover have th0: "a^n = prod (\<lambda>k. a) {0..m}" "(a oo c) ^ n = prod (\<lambda>k. a oo c) {0..m}"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3442
    by (simp_all add: prod_constant Suc)
67970
8c012a49293a A couple of new results
paulson <lp15@cam.ac.uk>
parents: 67411
diff changeset
  3443
  ultimately show ?thesis
8c012a49293a A couple of new results
paulson <lp15@cam.ac.uk>
parents: 67411
diff changeset
  3444
    by presburger
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3445
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3446
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3447
lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3448
  by (simp add: fps_eq_iff fps_compose_nth field_simps sum_negf[symmetric])
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3449
    
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3450
lemma fps_compose_sub_distrib: "(a - b) oo (c::'a::ring_1 fps) = (a oo c) - (b oo c)"
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
  3451
  using fps_compose_add_distrib [of a "- b" c] by (simp add: fps_compose_uminus)
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3452
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3453
lemma fps_X_fps_compose: "fps_X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3454
  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left sum.delta)
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3455
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3456
lemma fps_inverse_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3457
  assumes b0: "(b$0 :: 'a::field) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3458
    and a0: "a$0 \<noteq> 0"
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3459
  shows "inverse a oo b = inverse (a oo b)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3460
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3461
  let ?ia = "inverse a"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3462
  let ?ab = "a oo b"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3463
  let ?iab = "inverse ?ab"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3464
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3465
  from a0 have ia0: "?ia $ 0 \<noteq> 0" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3466
  from a0 have ab0: "?ab $ 0 \<noteq> 0" by (simp add: fps_compose_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3467
  have "(?ia oo b) *  (a oo b) = 1"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3468
    unfolding fps_compose_mult_distrib[OF b0, symmetric]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3469
    unfolding inverse_mult_eq_1[OF a0]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3470
    fps_compose_1 ..
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3471
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3472
  then have "(?ia oo b) *  (a oo b) * ?iab  = 1 * ?iab" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3473
  then have "(?ia oo b) *  (?iab * (a oo b))  = ?iab" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3474
  then show ?thesis unfolding inverse_mult_eq_1[OF ab0] by simp
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3475
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3476
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3477
lemma fps_divide_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3478
  assumes c0: "(c$0 :: 'a::field) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3479
    and b0: "b$0 \<noteq> 0"
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3480
  shows "(a/b) oo c = (a oo c) / (b oo c)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3481
    using b0 c0 by (simp add: fps_divide_unit fps_inverse_compose fps_compose_mult_distrib)
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3482
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3483
lemma gp:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3484
  assumes a0: "a$0 = (0::'a::field)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3485
  shows "(Abs_fps (\<lambda>n. 1)) oo a = 1/(1 - a)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3486
    (is "?one oo a = _")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3487
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3488
  have o0: "?one $ 0 \<noteq> 0" by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3489
  have th0: "(1 - fps_X) $ 0 \<noteq> (0::'a)" by simp
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3490
  from fps_inverse_gp[where ?'a = 'a]
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3491
  have "inverse ?one = 1 - fps_X" by (simp add: fps_eq_iff)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3492
  then have "inverse (inverse ?one) = inverse (1 - fps_X)" by simp
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3493
  then have th: "?one = 1/(1 - fps_X)" unfolding fps_inverse_idempotent[OF o0]
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3494
    by (simp add: fps_divide_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3495
  show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3496
    unfolding th
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3497
    unfolding fps_divide_compose[OF a0 th0]
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3498
    fps_compose_1 fps_compose_sub_distrib fps_X_fps_compose_startby0[OF a0] ..
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3499
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3500
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3501
lemma fps_const_power [simp]: "fps_const (c::'a::ring_1) ^ n = fps_const (c^n)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3502
  by (induct n) auto
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3503
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3504
lemma fps_compose_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3505
  assumes b0: "b$0 = (0::'a::field_char_0)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3506
    and ra0: "r (Suc k) (a$0) ^ Suc k = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3507
    and a0: "a$0 \<noteq> 0"
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3508
  shows "fps_radical r (Suc k)  a oo b = fps_radical r (Suc k) (a oo b)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3509
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3510
  let ?r = "fps_radical r (Suc k)"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3511
  let ?ab = "a oo b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3512
  have ab0: "?ab $ 0 = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3513
    by (simp add: fps_compose_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3514
  from ab0 a0 ra0 have rab0: "?ab $ 0 \<noteq> 0" "r (Suc k) (?ab $ 0) ^ Suc k = ?ab $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3515
    by simp_all
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3516
  have th00: "r (Suc k) ((a oo b) $ 0) = (fps_radical r (Suc k) a oo b) $ 0"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3517
    by (simp add: ab0 fps_compose_def)
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3518
  have th0: "(?r a oo b) ^ (Suc k) = a  oo b"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3519
    unfolding fps_compose_power[OF b0]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3520
    unfolding iffD1[OF power_radical[of a r k], OF a0 ra0]  ..
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3521
  from iffD1[OF radical_unique[where r=r and k=k and b= ?ab and a = "?r a oo b", OF rab0(2) th00 rab0(1)], OF th0]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3522
  show ?thesis  .
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3523
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3524
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3525
lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3526
  by (simp add: fps_eq_iff fps_compose_nth sum_distrib_left mult.assoc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3527
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3528
lemma fps_const_mult_apply_right:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3529
  "(a oo b) * fps_const (c::'a::comm_semiring_1) = (fps_const c * a) oo b"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  3530
  by (auto simp add: fps_const_mult_apply_left mult.commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3531
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3532
lemma fps_compose_assoc:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3533
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3534
    and b0: "b$0 = 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3535
  shows "a oo (b oo c) = a oo b oo c" (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3536
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3537
  have "?l$n = ?r$n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3538
  proof -
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3539
    have "?l$n = (sum (\<lambda>i. (fps_const (a$i) * b^i) oo c) {0..n})$n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3540
      by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3541
        sum_distrib_left mult.assoc fps_sum_nth)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3542
    also have "\<dots> = ((sum (\<lambda>i. fps_const (a$i) * b^i) {0..n}) oo c)$n"
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3543
      by (simp add: fps_compose_sum_distrib)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3544
    also have "\<dots> = ?r$n"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3545
      apply (simp add: fps_compose_nth fps_sum_nth sum_distrib_right mult.assoc)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3546
      apply (rule sum.cong)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3547
      apply (rule refl)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3548
      apply (rule sum.mono_neutral_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3549
      apply (auto simp add: not_le)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3550
      apply (erule startsby_zero_power_prefix[OF b0, rule_format])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3551
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3552
    finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3553
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3554
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3555
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3556
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3557
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3558
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3559
lemma fps_X_power_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3560
  assumes a0: "a$0=0"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3561
  shows "fps_X^k oo a = (a::'a::idom fps)^k"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3562
  (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3563
proof (cases k)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3564
  case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3565
  then show ?thesis by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3566
next
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3567
  case (Suc h)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3568
  have "?l $ n = ?r $n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3569
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3570
    consider "k > n" | "k \<le> n" by arith
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3571
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3572
    proof cases
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3573
      case 1
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3574
      then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3575
        using a0 startsby_zero_power_prefix[OF a0] Suc
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3576
        by (simp add: fps_compose_nth del: power_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3577
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3578
      case 2
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3579
      then show ?thesis
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3580
        by (simp add: fps_compose_nth mult_delta_left sum.delta)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3581
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3582
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3583
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3584
    unfolding fps_eq_iff by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3585
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3586
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3587
lemma fps_inv_right:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3588
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3589
    and a1: "a$1 \<noteq> 0"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3590
  shows "a oo fps_inv a = fps_X"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3591
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3592
  let ?ia = "fps_inv a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3593
  let ?iaa = "a oo fps_inv a"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3594
  have th0: "?ia $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3595
    by (simp add: fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3596
  have th1: "?iaa $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3597
    using a0 a1 by (simp add: fps_inv_def fps_compose_nth)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3598
  have th2: "fps_X$0 = 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3599
    by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3600
  from fps_inv[OF a0 a1] have "a oo (fps_inv a oo a) = a oo fps_X"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3601
    by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3602
  then have "(a oo fps_inv a) oo a = fps_X oo a"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3603
    by (simp add: fps_compose_assoc[OF a0 th0] fps_X_fps_compose_startby0[OF a0])
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3604
  with fps_compose_inj_right[OF a0 a1] show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3605
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3606
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3607
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3608
lemma fps_inv_deriv:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3609
  assumes a0: "a$0 = (0::'a::field)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3610
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3611
  shows "fps_deriv (fps_inv a) = inverse (fps_deriv a oo fps_inv a)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3612
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3613
  let ?ia = "fps_inv a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3614
  let ?d = "fps_deriv a oo ?ia"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3615
  let ?dia = "fps_deriv ?ia"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3616
  have ia0: "?ia$0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3617
    by (simp add: fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3618
  have th0: "?d$0 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3619
    using a1 by (simp add: fps_compose_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3620
  from fps_inv_right[OF a0 a1] have "?d * ?dia = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3621
    by (simp add: fps_compose_deriv[OF ia0, of a, symmetric] )
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3622
  then have "inverse ?d * ?d * ?dia = inverse ?d * 1"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3623
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3624
  with inverse_mult_eq_1 [OF th0] show "?dia = inverse ?d"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3625
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3626
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3627
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3628
lemma fps_inv_idempotent:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3629
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3630
    and a1: "a$1 \<noteq> 0"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3631
  shows "fps_inv (fps_inv a) = a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3632
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3633
  let ?r = "fps_inv"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3634
  have ra0: "?r a $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3635
    by (simp add: fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3636
  from a1 have ra1: "?r a $ 1 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3637
    by (simp add: fps_inv_def field_simps)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3638
  have fps_X0: "fps_X$0 = 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3639
    by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3640
  from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = fps_X" .
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3641
  then have "?r (?r a) oo ?r a oo a = fps_X oo a"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3642
    by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3643
  then have "?r (?r a) oo (?r a oo a) = a"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3644
    unfolding fps_X_fps_compose_startby0[OF a0]
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3645
    unfolding fps_compose_assoc[OF a0 ra0, symmetric] .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3646
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3647
    unfolding fps_inv[OF a0 a1] by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3648
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3649
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3650
lemma fps_ginv_ginv:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3651
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3652
    and a1: "a$1 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3653
    and c0: "c$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3654
    and  c1: "c$1 \<noteq> 0"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3655
  shows "fps_ginv b (fps_ginv c a) = b oo a oo fps_inv c"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3656
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3657
  let ?r = "fps_ginv"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3658
  from c0 have rca0: "?r c a $0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3659
    by (simp add: fps_ginv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3660
  from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3661
    by (simp add: fps_ginv_def field_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3662
  from fps_ginv[OF rca0 rca1]
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3663
  have "?r b (?r c a) oo ?r c a = b" .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3664
  then have "?r b (?r c a) oo ?r c a oo a = b oo a"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3665
    by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3666
  then have "?r b (?r c a) oo (?r c a oo a) = b oo a"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3667
    apply (subst fps_compose_assoc)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3668
    using a0 c0
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3669
    apply (auto simp add: fps_ginv_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3670
    done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3671
  then have "?r b (?r c a) oo c = b oo a"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3672
    unfolding fps_ginv[OF a0 a1] .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3673
  then have "?r b (?r c a) oo c oo fps_inv c= b oo a oo fps_inv c"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3674
    by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3675
  then have "?r b (?r c a) oo (c oo fps_inv c) = b oo a oo fps_inv c"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3676
    apply (subst fps_compose_assoc)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3677
    using a0 c0
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3678
    apply (auto simp add: fps_inv_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3679
    done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3680
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3681
    unfolding fps_inv_right[OF c0 c1] by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3682
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3683
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3684
lemma fps_ginv_deriv:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3685
  assumes a0:"a$0 = (0::'a::field)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3686
    and a1: "a$1 \<noteq> 0"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3687
  shows "fps_deriv (fps_ginv b a) = (fps_deriv b / fps_deriv a) oo fps_ginv fps_X a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3688
proof -
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3689
  let ?ia = "fps_ginv b a"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3690
  let ?ifps_Xa = "fps_ginv fps_X a"
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3691
  let ?d = "fps_deriv"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3692
  let ?dia = "?d ?ia"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3693
  have ifps_Xa0: "?ifps_Xa $ 0 = 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3694
    by (simp add: fps_ginv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3695
  have da0: "?d a $ 0 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3696
    using a1 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3697
  from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3698
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3699
  then have "(?d ?ia oo a) * ?d a = ?d b"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3700
    unfolding fps_compose_deriv[OF a0] .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3701
  then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3702
    by simp
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3703
  with a1 have "(?d ?ia oo a) * (inverse (?d a) * ?d a) = ?d b / ?d a"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3704
    by (simp add: fps_divide_unit)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3705
  then have "(?d ?ia oo a) oo ?ifps_Xa =  (?d b / ?d a) oo ?ifps_Xa"
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3706
    unfolding inverse_mult_eq_1[OF da0] by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3707
  then have "?d ?ia oo (a oo ?ifps_Xa) =  (?d b / ?d a) oo ?ifps_Xa"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3708
    unfolding fps_compose_assoc[OF ifps_Xa0 a0] .
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3709
  then show ?thesis unfolding fps_inv_ginv[symmetric]
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3710
    unfolding fps_inv_right[OF a0 a1] by simp
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3711
qed
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3712
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3713
lemma fps_compose_linear:
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3714
  "fps_compose (f :: 'a :: comm_ring_1 fps) (fps_const c * fps_X) = Abs_fps (\<lambda>n. c^n * f $ n)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  3715
  by (simp add: fps_eq_iff fps_compose_def power_mult_distrib
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  3716
                if_distrib sum.delta' cong: if_cong)
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3717
              
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3718
lemma fps_compose_uminus': 
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3719
  "fps_compose f (-fps_X :: 'a :: comm_ring_1 fps) = Abs_fps (\<lambda>n. (-1)^n * f $ n)"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3720
  using fps_compose_linear[of f "-1"] 
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3721
  by (simp only: fps_const_neg [symmetric] fps_const_1_eq_1) simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3722
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3723
subsection \<open>Elementary series\<close>
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3724
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3725
subsubsection \<open>Exponential series\<close>
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3726
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3727
definition "fps_exp x = Abs_fps (\<lambda>n. x^n / of_nat (fact n))"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3728
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3729
lemma fps_exp_deriv[simp]: "fps_deriv (fps_exp a) = fps_const (a::'a::field_char_0) * fps_exp a" 
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3730
  (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3731
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3732
  have "?l$n = ?r $ n" for n
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3733
    apply (auto simp add: fps_exp_def field_simps power_Suc[symmetric]
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63317
diff changeset
  3734
      simp del: fact_Suc of_nat_Suc power_Suc)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  3735
    apply (simp add: field_simps)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3736
    done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3737
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3738
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3739
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3740
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3741
lemma fps_exp_unique_ODE:
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3742
  "fps_deriv a = fps_const c * a \<longleftrightarrow> a = fps_const (a$0) * fps_exp (c::'a::field_char_0)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3743
  (is "?lhs \<longleftrightarrow> ?rhs")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3744
proof
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3745
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3746
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3747
    from that have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3748
      by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3749
    have th': "a$n = a$0 * c ^ n/ (fact n)" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3750
    proof (induct n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3751
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3752
      then show ?case by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3753
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3754
      case Suc
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3755
      then show ?case
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3756
        unfolding th
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3757
        using fact_gt_zero
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3758
        apply (simp add: field_simps del: of_nat_Suc fact_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3759
        apply simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3760
        done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3761
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3762
    show ?thesis
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3763
      by (auto simp add: fps_eq_iff fps_const_mult_left fps_exp_def intro: th')
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3764
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3765
  show ?lhs if ?rhs
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3766
    using that by (metis fps_exp_deriv fps_deriv_mult_const_left mult.left_commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3767
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3768
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3769
lemma fps_exp_add_mult: "fps_exp (a + b) = fps_exp (a::'a::field_char_0) * fps_exp b" (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3770
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3771
  have "fps_deriv ?r = fps_const (a + b) * ?r"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3772
    by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3773
  then have "?r = ?l"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3774
    by (simp only: fps_exp_unique_ODE) (simp add: fps_mult_nth fps_exp_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3775
  then show ?thesis ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3776
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3777
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3778
lemma fps_exp_nth[simp]: "fps_exp a $ n = a^n / of_nat (fact n)"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3779
  by (simp add: fps_exp_def)
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3780
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3781
lemma fps_exp_0[simp]: "fps_exp (0::'a::field) = 1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3782
  by (simp add: fps_eq_iff power_0_left)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3783
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3784
lemma fps_exp_neg: "fps_exp (- a) = inverse (fps_exp (a::'a::field_char_0))"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3785
proof -
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3786
  from fps_exp_add_mult[of a "- a"] have th0: "fps_exp a * fps_exp (- a) = 1" by simp
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3787
  from fps_inverse_unique[OF th0] show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3788
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3789
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3790
lemma fps_exp_nth_deriv[simp]: 
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3791
  "fps_nth_deriv n (fps_exp (a::'a::field_char_0)) = (fps_const a)^n * (fps_exp a)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3792
  by (induct n) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3793
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3794
lemma fps_X_compose_fps_exp[simp]: "fps_X oo fps_exp (a::'a::field) = fps_exp a - 1"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3795
  by (simp add: fps_eq_iff fps_X_fps_compose)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3796
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3797
lemma fps_inv_fps_exp_compose:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3798
  assumes a: "a \<noteq> 0"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3799
  shows "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = fps_X"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3800
    and "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_X"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3801
proof -
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3802
  let ?b = "fps_exp a - 1"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3803
  have b0: "?b $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3804
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3805
  have b1: "?b $ 1 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3806
    by (simp add: a)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3807
  from fps_inv[OF b0 b1] show "fps_inv (fps_exp a - 1) oo (fps_exp a - 1) = fps_X" .
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3808
  from fps_inv_right[OF b0 b1] show "(fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_X" .
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3809
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3810
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3811
lemma fps_exp_power_mult: "(fps_exp (c::'a::field_char_0))^n = fps_exp (of_nat n * c)"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3812
  by (induct n) (auto simp add: field_simps fps_exp_add_mult)
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3813
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3814
lemma radical_fps_exp:
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3815
  assumes r: "r (Suc k) 1 = 1"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3816
  shows "fps_radical r (Suc k) (fps_exp (c::'a::field_char_0)) = fps_exp (c / of_nat (Suc k))"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3817
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3818
  let ?ck = "(c / of_nat (Suc k))"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3819
  let ?r = "fps_radical r (Suc k)"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3820
  have eq0[simp]: "?ck * of_nat (Suc k) = c" "of_nat (Suc k) * ?ck = c"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3821
    by (simp_all del: of_nat_Suc)
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3822
  have th0: "fps_exp ?ck ^ (Suc k) = fps_exp c" unfolding fps_exp_power_mult eq0 ..
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3823
  have th: "r (Suc k) (fps_exp c $0) ^ Suc k = fps_exp c $ 0"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3824
    "r (Suc k) (fps_exp c $ 0) = fps_exp ?ck $ 0" "fps_exp c $ 0 \<noteq> 0" using r by simp_all
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3825
  from th0 radical_unique[where r=r and k=k, OF th] show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3826
    by auto
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3827
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3828
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3829
lemma fps_exp_compose_linear [simp]: 
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3830
  "fps_exp (d::'a::field_char_0) oo (fps_const c * fps_X) = fps_exp (c * d)"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3831
  by (simp add: fps_compose_linear fps_exp_def fps_eq_iff power_mult_distrib)
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3832
  
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3833
lemma fps_fps_exp_compose_minus [simp]: 
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3834
  "fps_compose (fps_exp c) (-fps_X) = fps_exp (-c :: 'a :: field_char_0)"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3835
  using fps_exp_compose_linear[of c "-1 :: 'a"] 
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3836
  unfolding fps_const_neg [symmetric] fps_const_1_eq_1 by simp
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3837
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3838
lemma fps_exp_eq_iff [simp]: "fps_exp c = fps_exp d \<longleftrightarrow> c = (d :: 'a :: field_char_0)"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3839
proof
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3840
  assume "fps_exp c = fps_exp d"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3841
  from arg_cong[of _ _ "\<lambda>F. F $ 1", OF this] show "c = d" by simp
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3842
qed simp_all
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3843
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3844
lemma fps_exp_eq_fps_const_iff [simp]: 
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3845
  "fps_exp (c :: 'a :: field_char_0) = fps_const c' \<longleftrightarrow> c = 0 \<and> c' = 1"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3846
proof
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3847
  assume "c = 0 \<and> c' = 1"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3848
  thus "fps_exp c = fps_const c'" by (auto simp: fps_eq_iff)
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3849
next
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3850
  assume "fps_exp c = fps_const c'"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3851
  from arg_cong[of _ _ "\<lambda>F. F $ 1", OF this] arg_cong[of _ _ "\<lambda>F. F $ 0", OF this] 
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3852
    show "c = 0 \<and> c' = 1" by simp_all
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3853
qed
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3854
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3855
lemma fps_exp_neq_0 [simp]: "\<not>fps_exp (c :: 'a :: field_char_0) = 0"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3856
  unfolding fps_const_0_eq_0 [symmetric] fps_exp_eq_fps_const_iff by simp  
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3857
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3858
lemma fps_exp_eq_1_iff [simp]: "fps_exp (c :: 'a :: field_char_0) = 1 \<longleftrightarrow> c = 0"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3859
  unfolding fps_const_1_eq_1 [symmetric] fps_exp_eq_fps_const_iff by simp
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3860
    
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3861
lemma fps_exp_neq_numeral_iff [simp]: 
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3862
  "fps_exp (c :: 'a :: field_char_0) = numeral n \<longleftrightarrow> c = 0 \<and> n = Num.One"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3863
  unfolding numeral_fps_const fps_exp_eq_fps_const_iff by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3864
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3865
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3866
subsubsection \<open>Logarithmic series\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3867
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3868
lemma Abs_fps_if_0:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3869
  "Abs_fps (\<lambda>n. if n = 0 then (v::'a::ring_1) else f n) =
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3870
    fps_const v + fps_X * Abs_fps (\<lambda>n. f (Suc n))"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3871
  by (auto simp add: fps_eq_iff)
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3872
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3873
definition fps_ln :: "'a::field_char_0 \<Rightarrow> 'a fps"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3874
  where "fps_ln c = fps_const (1/c) * Abs_fps (\<lambda>n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n)"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3875
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3876
lemma fps_ln_deriv: "fps_deriv (fps_ln c) = fps_const (1/c) * inverse (1 + fps_X)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3877
  unfolding fps_inverse_fps_X_plus1
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3878
  by (simp add: fps_ln_def fps_eq_iff del: of_nat_Suc)
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3879
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3880
lemma fps_ln_nth: "fps_ln c $ n = (if n = 0 then 0 else 1/c * ((- 1) ^ (n - 1) / of_nat n))"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3881
  by (simp add: fps_ln_def field_simps)
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3882
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3883
lemma fps_ln_0 [simp]: "fps_ln c $ 0 = 0" by (simp add: fps_ln_def)
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3884
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3885
lemma fps_ln_fps_exp_inv:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3886
  fixes a :: "'a::field_char_0"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3887
  assumes a: "a \<noteq> 0"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3888
  shows "fps_ln a = fps_inv (fps_exp a - 1)"  (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3889
proof -
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3890
  let ?b = "fps_exp a - 1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3891
  have b0: "?b $ 0 = 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3892
  have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3893
  have "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) =
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3894
    (fps_const a * (fps_exp a - 1) + fps_const a) oo fps_inv (fps_exp a - 1)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3895
    by (simp add: field_simps)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3896
  also have "\<dots> = fps_const a * (fps_X + 1)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3897
    apply (simp add: fps_compose_add_distrib fps_const_mult_apply_left[symmetric] fps_inv_right[OF b0 b1])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3898
    apply (simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3899
    done
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3900
  finally have eq: "fps_deriv (fps_exp a - 1) oo fps_inv (fps_exp a - 1) = fps_const a * (fps_X + 1)" .
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3901
  from fps_inv_deriv[OF b0 b1, unfolded eq]
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3902
  have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (fps_X + 1)"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3903
    using a by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3904
  then have "fps_deriv ?l = fps_deriv ?r"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3905
    by (simp add: fps_ln_deriv add.commute fps_divide_def divide_inverse)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3906
  then show ?thesis unfolding fps_deriv_eq_iff
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3907
    by (simp add: fps_ln_nth fps_inv_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3908
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3909
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3910
lemma fps_ln_mult_add:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3911
  assumes c0: "c\<noteq>0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3912
    and d0: "d\<noteq>0"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3913
  shows "fps_ln c + fps_ln d = fps_const (c+d) * fps_ln (c*d)"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3914
  (is "?r = ?l")
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3915
proof-
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3916
  from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_simps)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3917
  have "fps_deriv ?r = fps_const (1/c + 1/d) * inverse (1 + fps_X)"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3918
    by (simp add: fps_ln_deriv fps_const_add[symmetric] algebra_simps del: fps_const_add)
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3919
  also have "\<dots> = fps_deriv ?l"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3920
    apply (simp add: fps_ln_deriv)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3921
    apply (simp add: fps_eq_iff eq)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3922
    done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3923
  finally show ?thesis
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3924
    unfolding fps_deriv_eq_iff by simp
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3925
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3926
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3927
lemma fps_X_dvd_fps_ln [simp]: "fps_X dvd fps_ln c"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3928
proof -
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3929
  have "fps_ln c = fps_X * Abs_fps (\<lambda>n. (-1) ^ n / (of_nat (Suc n) * c))"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3930
    by (intro fps_ext) (auto simp: fps_ln_def of_nat_diff)
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3931
  thus ?thesis by simp
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3932
qed
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  3933
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3934
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3935
subsubsection \<open>Binomial series\<close>
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3936
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3937
definition "fps_binomial a = Abs_fps (\<lambda>n. a gchoose n)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3938
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3939
lemma fps_binomial_nth[simp]: "fps_binomial a $ n = a gchoose n"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3940
  by (simp add: fps_binomial_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3941
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3942
lemma fps_binomial_ODE_unique:
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3943
  fixes c :: "'a::field_char_0"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3944
  shows "fps_deriv a = (fps_const c * a) / (1 + fps_X) \<longleftrightarrow> a = fps_const (a$0) * fps_binomial c"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3945
  (is "?lhs \<longleftrightarrow> ?rhs")
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3946
proof
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3947
  let ?da = "fps_deriv a"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  3948
  let ?x1 = "(1 + fps_X):: 'a fps"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3949
  let ?l = "?x1 * ?da"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3950
  let ?r = "fps_const c * a"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3951
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3952
  have eq: "?l = ?r \<longleftrightarrow> ?lhs"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3953
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3954
    have x10: "?x1 $ 0 \<noteq> 0" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3955
    have "?l = ?r \<longleftrightarrow> inverse ?x1 * ?l = inverse ?x1 * ?r" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3956
    also have "\<dots> \<longleftrightarrow> ?da = (fps_const c * a) / ?x1"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3957
      apply (simp only: fps_divide_def  mult.assoc[symmetric] inverse_mult_eq_1[OF x10])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3958
      apply (simp add: field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3959
      done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3960
    finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3961
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3962
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3963
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3964
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3965
    from eq that have h: "?l = ?r" ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3966
    have th0: "a$ Suc n = ((c - of_nat n) / of_nat (Suc n)) * a $n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3967
    proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3968
      from h have "?l $ n = ?r $ n" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3969
      then show ?thesis
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3970
        apply (simp add: field_simps del: of_nat_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3971
        apply (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3972
        apply (simp_all add: field_simps del: of_nat_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3973
        done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3974
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3975
    have th1: "a $ n = (c gchoose n) * a $ 0" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3976
    proof (induct n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3977
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3978
      then show ?case by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3979
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3980
      case (Suc m)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3981
      then show ?case
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3982
        unfolding th0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3983
        apply (simp add: field_simps del: of_nat_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3984
        unfolding mult.assoc[symmetric] gbinomial_mult_1
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3985
        apply (simp add: field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3986
        done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3987
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3988
    show ?thesis
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3989
      apply (simp add: fps_eq_iff)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3990
      apply (subst th1)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3991
      apply (simp add: field_simps)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3992
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3993
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3994
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3995
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3996
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3997
    have th00: "x * (a $ 0 * y) = a $ 0 * (x * y)" for x y
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  3998
      by (simp add: mult.commute)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3999
    have "?l = ?r"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4000
      apply (subst \<open>?rhs\<close>)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4001
      apply (subst (2) \<open>?rhs\<close>)
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  4002
      apply (clarsimp simp add: fps_eq_iff field_simps)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  4003
      unfolding mult.assoc[symmetric] th00 gbinomial_mult_1
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4004
      apply (simp add: field_simps gbinomial_mult_1)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4005
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4006
    with eq show ?thesis ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4007
  qed
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4008
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4009
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4010
lemma fps_binomial_ODE_unique':
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4011
  "(fps_deriv a = fps_const c * a / (1 + fps_X) \<and> a $ 0 = 1) \<longleftrightarrow> (a = fps_binomial c)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4012
  by (subst fps_binomial_ODE_unique) auto
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4013
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4014
lemma fps_binomial_deriv: "fps_deriv (fps_binomial c) = fps_const c * fps_binomial c / (1 + fps_X)"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4015
proof -
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4016
  let ?a = "fps_binomial c"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4017
  have th0: "?a = fps_const (?a$0) * ?a" by (simp)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4018
  from iffD2[OF fps_binomial_ODE_unique, OF th0] show ?thesis .
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4019
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4020
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4021
lemma fps_binomial_add_mult: "fps_binomial (c+d) = fps_binomial c * fps_binomial d" (is "?l = ?r")
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4022
proof -
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4023
  let ?P = "?r - ?l"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4024
  let ?b = "fps_binomial"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4025
  let ?db = "\<lambda>x. fps_deriv (?b x)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4026
  have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)"  by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4027
  also have "\<dots> = inverse (1 + fps_X) *
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4028
      (fps_const c * ?b c * ?b d + fps_const d * ?b c * ?b d - fps_const (c+d) * ?b (c + d))"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4029
    unfolding fps_binomial_deriv
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  4030
    by (simp add: fps_divide_def field_simps)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4031
  also have "\<dots> = (fps_const (c + d)/ (1 + fps_X)) * ?P"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4032
    by (simp add: field_simps fps_divide_unit fps_const_add[symmetric] del: fps_const_add)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4033
  finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + fps_X)"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4034
    by (simp add: fps_divide_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4035
  have "?P = fps_const (?P$0) * ?b (c + d)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4036
    unfolding fps_binomial_ODE_unique[symmetric]
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4037
    using th0 by simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4038
  then have "?P = 0" by (simp add: fps_mult_nth)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4039
  then show ?thesis by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4040
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4041
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4042
lemma fps_binomial_minus_one: "fps_binomial (- 1) = inverse (1 + fps_X)"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4043
  (is "?l = inverse ?r")
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4044
proof-
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4045
  have th: "?r$0 \<noteq> 0" by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4046
  have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + fps_X)"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4047
    by (simp add: fps_inverse_deriv[OF th] fps_divide_def
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  4048
      power2_eq_square mult.commute fps_const_neg[symmetric] del: fps_const_neg)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4049
  have eq: "inverse ?r $ 0 = 1"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4050
    by (simp add: fps_inverse_def)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4051
  from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + fps_X)" "- 1"] th'] eq
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4052
  show ?thesis by (simp add: fps_inverse_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4053
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4054
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4055
lemma fps_binomial_of_nat: "fps_binomial (of_nat n) = (1 + fps_X :: 'a :: field_char_0 fps) ^ n"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4056
proof (cases "n = 0")
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4057
  case [simp]: True
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4058
  have "fps_deriv ((1 + fps_X) ^ n :: 'a fps) = 0" by simp
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4059
  also have "\<dots> = fps_const (of_nat n) * (1 + fps_X) ^ n / (1 + fps_X)" by (simp add: fps_binomial_def)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4060
  finally show ?thesis by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) simp_all
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4061
next
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4062
  case False
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4063
  have "fps_deriv ((1 + fps_X) ^ n :: 'a fps) = fps_const (of_nat n) * (1 + fps_X) ^ (n - 1)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4064
    by (simp add: fps_deriv_power)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4065
  also have "(1 + fps_X :: 'a fps) $ 0 \<noteq> 0" by simp
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4066
  hence "(1 + fps_X :: 'a fps) \<noteq> 0" by (intro notI) (simp only: , simp)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4067
  with False have "(1 + fps_X :: 'a fps) ^ (n - 1) = (1 + fps_X) ^ n / (1 + fps_X)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4068
    by (cases n) (simp_all )
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4069
  also have "fps_const (of_nat n :: 'a) * ((1 + fps_X) ^ n / (1 + fps_X)) =
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4070
               fps_const (of_nat n) * (1 + fps_X) ^ n / (1 + fps_X)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4071
    by (simp add: unit_div_mult_swap)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4072
  finally show ?thesis
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4073
    by (subst sym, subst fps_binomial_ODE_unique' [symmetric]) (simp_all add: fps_power_nth)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4074
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4075
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4076
lemma fps_binomial_0 [simp]: "fps_binomial 0 = 1"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4077
  using fps_binomial_of_nat[of 0] by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4078
  
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4079
lemma fps_binomial_power: "fps_binomial a ^ n = fps_binomial (of_nat n * a)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4080
  by (induction n) (simp_all add: fps_binomial_add_mult ring_distribs)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4081
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4082
lemma fps_binomial_1: "fps_binomial 1 = 1 + fps_X"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4083
  using fps_binomial_of_nat[of 1] by simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4084
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4085
lemma fps_binomial_minus_of_nat:
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4086
  "fps_binomial (- of_nat n) = inverse ((1 + fps_X :: 'a :: field_char_0 fps) ^ n)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4087
  by (rule sym, rule fps_inverse_unique)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4088
     (simp add: fps_binomial_of_nat [symmetric] fps_binomial_add_mult [symmetric])
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4089
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4090
lemma one_minus_const_fps_X_power:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4091
  "c \<noteq> 0 \<Longrightarrow> (1 - fps_const c * fps_X) ^ n =
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4092
     fps_compose (fps_binomial (of_nat n)) (-fps_const c * fps_X)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4093
  by (subst fps_binomial_of_nat)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4094
     (simp add: fps_compose_power [symmetric] fps_compose_add_distrib fps_const_neg [symmetric] 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4095
           del: fps_const_neg)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4096
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4097
lemma one_minus_fps_X_const_neg_power:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4098
  "inverse ((1 - fps_const c * fps_X) ^ n) = 
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4099
       fps_compose (fps_binomial (-of_nat n)) (-fps_const c * fps_X)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4100
proof (cases "c = 0")
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4101
  case False
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4102
  thus ?thesis
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4103
  by (subst fps_binomial_minus_of_nat)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4104
     (simp add: fps_compose_power [symmetric] fps_inverse_compose fps_compose_add_distrib
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4105
                fps_const_neg [symmetric] del: fps_const_neg)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4106
qed simp
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4107
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4108
lemma fps_X_plus_const_power:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4109
  "c \<noteq> 0 \<Longrightarrow> (fps_X + fps_const c) ^ n =
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4110
     fps_const (c^n) * fps_compose (fps_binomial (of_nat n)) (fps_const (inverse c) * fps_X)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4111
  by (subst fps_binomial_of_nat)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4112
     (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4113
                fps_const_power [symmetric] power_mult_distrib [symmetric] 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4114
                algebra_simps inverse_mult_eq_1' del: fps_const_power)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4115
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4116
lemma fps_X_plus_const_neg_power:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4117
  "c \<noteq> 0 \<Longrightarrow> inverse ((fps_X + fps_const c) ^ n) =
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4118
     fps_const (inverse c^n) * fps_compose (fps_binomial (-of_nat n)) (fps_const (inverse c) * fps_X)"
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4119
  by (subst fps_binomial_minus_of_nat)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4120
     (simp add: fps_compose_power [symmetric] fps_binomial_of_nat fps_compose_add_distrib
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4121
                fps_const_power [symmetric] power_mult_distrib [symmetric] fps_inverse_compose 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4122
                algebra_simps fps_const_inverse [symmetric] fps_inverse_mult [symmetric]
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4123
                fps_inverse_power [symmetric] inverse_mult_eq_1'
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4124
           del: fps_const_power)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4125
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4126
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4127
lemma one_minus_const_fps_X_neg_power':
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4128
  "n > 0 \<Longrightarrow> inverse ((1 - fps_const (c :: 'a :: field_char_0) * fps_X) ^ n) =
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4129
       Abs_fps (\<lambda>k. of_nat ((n + k - 1) choose k) * c^k)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4130
  apply (rule fps_ext)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4131
  apply (subst one_minus_fps_X_const_neg_power, subst fps_const_neg, subst fps_compose_linear)
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4132
  apply (simp add: power_mult_distrib [symmetric] mult.assoc [symmetric] 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4133
                   gbinomial_minus binomial_gbinomial of_nat_diff)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4134
  done
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4135
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4136
text \<open>Vandermonde's Identity as a consequence.\<close>
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4137
lemma gbinomial_Vandermonde:
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4138
  "sum (\<lambda>k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4139
proof -
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4140
  let ?ba = "fps_binomial a"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4141
  let ?bb = "fps_binomial b"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4142
  let ?bab = "fps_binomial (a + b)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4143
  from fps_binomial_add_mult[of a b] have "?bab $ n = (?ba * ?bb)$n" by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4144
  then show ?thesis by (simp add: fps_mult_nth)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4145
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4146
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4147
lemma binomial_Vandermonde:
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4148
  "sum (\<lambda>k. (a choose k) * (b choose (n - k))) {0..n} = (a + b) choose n"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4149
  using gbinomial_Vandermonde[of "(of_nat a)" "of_nat b" n]
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
  4150
  by (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4151
                 of_nat_sum[symmetric] of_nat_add[symmetric] of_nat_eq_iff)
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4152
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4153
lemma binomial_Vandermonde_same: "sum (\<lambda>k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4154
  using binomial_Vandermonde[of n n n, symmetric]
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4155
  unfolding mult_2
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4156
  apply (simp add: power2_eq_square)
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4157
  apply (rule sum.cong)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4158
  apply (auto intro:  binomial_symmetric)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4159
  done
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4160
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4161
lemma Vandermonde_pochhammer_lemma:
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4162
  fixes a :: "'a::field_char_0"
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4163
  assumes b: "\<forall>j\<in>{0 ..<n}. b \<noteq> of_nat j"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4164
  shows "sum (\<lambda>k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) /
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4165
      (of_nat (fact k) * pochhammer (b - of_nat n + 1) k)) {0..n} =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4166
    pochhammer (- (a + b)) n / pochhammer (- b) n"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4167
  (is "?l = ?r")
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4168
proof -
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4169
  let ?m1 = "\<lambda>m. (- 1 :: 'a) ^ m"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4170
  let ?f = "\<lambda>m. of_nat (fact m)"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4171
  let ?p = "\<lambda>(x::'a). pochhammer (- x)"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4172
  from b have bn0: "?p b n \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4173
    unfolding pochhammer_eq_0_iff by simp
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4174
  have th00:
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4175
    "b gchoose (n - k) =
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4176
        (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4177
      (is ?gchoose)
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4178
    "pochhammer (1 + b - of_nat n) k \<noteq> 0"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4179
      (is ?pochhammer)
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4180
    if kn: "k \<in> {0..n}" for k
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4181
  proof -
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4182
    from kn have "k \<le> n" by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4183
    have nz: "pochhammer (1 + b - of_nat n) n \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4184
    proof
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4185
      assume "pochhammer (1 + b - of_nat n) n = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4186
      then have c: "pochhammer (b - of_nat n + 1) n = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4187
        by (simp add: algebra_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4188
      then obtain j where j: "j < n" "b - of_nat n + 1 = - of_nat j"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4189
        unfolding pochhammer_eq_0_iff by blast
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4190
      from j have "b = of_nat n - of_nat j - of_nat 1"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4191
        by (simp add: algebra_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4192
      then have "b = of_nat (n - j - 1)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4193
        using j kn by (simp add: of_nat_diff)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4194
      with b show False using j by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4195
    qed
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4196
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4197
    from nz kn [simplified] have nz': "pochhammer (1 + b - of_nat n) k \<noteq> 0"
35175
61255c81da01 fix more looping simp rules
huffman
parents: 32960
diff changeset
  4198
      by (rule pochhammer_neq_0_mono)
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4199
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  4200
    consider "k = 0 \<or> n = 0" | "k \<noteq> 0" "n \<noteq> 0"
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  4201
      by blast
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4202
    then have "b gchoose (n - k) =
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4203
      (?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k)"
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4204
    proof cases
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4205
      case 1
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4206
      then show ?thesis
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4207
        using kn by (cases "k = 0") (simp_all add: gbinomial_pochhammer)
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4208
    next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  4209
      case neq: 2
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4210
      then obtain m where m: "n = Suc m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4211
        by (cases n) auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  4212
      from neq(1) obtain h where h: "k = Suc h"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4213
        by (cases k) auto
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4214
      show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4215
      proof (cases "k = n")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4216
        case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4217
        then show ?thesis
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59815
diff changeset
  4218
          using pochhammer_minus'[where k=k and b=b]
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59815
diff changeset
  4219
          apply (simp add: pochhammer_same)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4220
          using bn0
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4221
          apply (simp add: field_simps power_add[symmetric])
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4222
          done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4223
      next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4224
        case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4225
        with kn have kn': "k < n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4226
          by simp
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4227
        have m1nk: "?m1 n = prod (\<lambda>i. - 1) {..m}" "?m1 k = prod (\<lambda>i. - 1) {0..h}"
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4228
          by (simp_all add: prod_constant m h)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4229
        have bnz0: "pochhammer (b - of_nat n + 1) k \<noteq> 0"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4230
          using bn0 kn
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4231
          unfolding pochhammer_eq_0_iff
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4232
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4233
          apply (erule_tac x= "n - ka - 1" in allE)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4234
          apply (auto simp add: algebra_simps of_nat_diff)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4235
          done
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4236
        have eq1: "prod (\<lambda>k. (1::'a) + of_nat m - of_nat k) {..h} =
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4237
          prod of_nat {Suc (m - h) .. Suc m}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4238
          using kn' h m
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4239
          by (intro prod.reindex_bij_witness[where i="\<lambda>k. Suc m - k" and j="\<lambda>k. Suc m - k"])
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  4240
             (auto simp: of_nat_diff)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4241
        have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4242
          apply (simp add: pochhammer_minus field_simps)
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4243
          using \<open>k \<le> n\<close> apply (simp add: fact_split [of k n])
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4244
          apply (simp add: pochhammer_prod)
67411
3f4b0c84630f restored naming of lemmas after corresponding constants
haftmann
parents: 67399
diff changeset
  4245
          using prod.atLeastLessThan_shift_bounds [where ?'a = 'a, of "\<lambda>i. 1 + of_nat i" 0 "n - k" k]
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4246
          apply (auto simp add: of_nat_diff field_simps)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4247
          done
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4248
        have th20: "?m1 n * ?p b n = prod (\<lambda>i. b - of_nat i) {0..m}"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4249
          apply (simp add: pochhammer_minus field_simps m)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4250
          apply (auto simp add: pochhammer_prod_rev of_nat_diff prod.atLeast_Suc_atMost_Suc_shift)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4251
          done
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4252
        have th21:"pochhammer (b - of_nat n + 1) k = prod (\<lambda>i. b - of_nat i) {n - k .. n - 1}"
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4253
          using kn apply (simp add: pochhammer_prod_rev m h prod.atLeast_Suc_atMost_Suc_shift)
67411
3f4b0c84630f restored naming of lemmas after corresponding constants
haftmann
parents: 67399
diff changeset
  4254
          using prod.atLeastAtMost_shift_0 [of "m - h" m, where ?'a = 'a]
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4255
          apply (auto simp add: of_nat_diff field_simps)
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4256
          done
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4257
        have "?m1 n * ?p b n =
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4258
          prod (\<lambda>i. b - of_nat i) {0.. n - k - 1} * pochhammer (b - of_nat n + 1) k"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4259
          using kn' m h unfolding th20 th21 apply simp
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4260
          apply (subst prod.union_disjoint [symmetric])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4261
          apply auto
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4262
          apply (rule prod.cong)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4263
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4264
          done
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4265
        then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k =
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4266
          prod (\<lambda>i. b - of_nat i) {0.. n - k - 1}"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  4267
          using nz' by (simp add: field_simps)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4268
        have "(?m1 n * ?p b n * ?m1 k * ?p (of_nat n) k) / (?f n * pochhammer (b - of_nat n + 1) k) =
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4269
          ((?m1 k * ?p (of_nat n) k) / ?f n) * ((?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4270
          using bnz0
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  4271
          by (simp add: field_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4272
        also have "\<dots> = b gchoose (n - k)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  4273
          unfolding th1 th2
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4274
          using kn' m h
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4275
          apply (simp add: field_simps gbinomial_mult_fact)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  4276
          apply (rule prod.cong)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4277
          apply auto
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4278
          done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4279
        finally show ?thesis by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4280
      qed
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4281
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4282
    then show ?gchoose and ?pochhammer
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4283
      apply (cases "n = 0")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4284
      using nz'
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4285
      apply auto
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4286
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4287
  qed
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  4288
  have "?r = ((a + b) gchoose n) * (of_nat (fact n) / (?m1 n * pochhammer (- b) n))"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4289
    unfolding gbinomial_pochhammer
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  4290
    using bn0 by (auto simp add: field_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4291
  also have "\<dots> = ?l"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4292
    unfolding gbinomial_Vandermonde[symmetric]
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4293
    apply (simp add: th00)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4294
    unfolding gbinomial_pochhammer
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4295
    using bn0
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4296
    apply (simp add: sum_distrib_right sum_distrib_left field_simps)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4297
    done
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4298
  finally show ?thesis by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4299
qed
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4300
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4301
lemma Vandermonde_pochhammer:
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4302
  fixes a :: "'a::field_char_0"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4303
  assumes c: "\<forall>i \<in> {0..< n}. c \<noteq> - of_nat i"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4304
  shows "sum (\<lambda>k. (pochhammer a k * pochhammer (- (of_nat n)) k) /
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4305
    (of_nat (fact k) * pochhammer c k)) {0..n} = pochhammer (c - a) n / pochhammer c n"
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4306
proof -
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4307
  let ?a = "- a"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4308
  let ?b = "c + of_nat n - 1"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4309
  have h: "\<forall> j \<in>{0..< n}. ?b \<noteq> of_nat j"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4310
    using c
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4311
    apply (auto simp add: algebra_simps of_nat_diff)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4312
    apply (erule_tac x = "n - j - 1" in ballE)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4313
    apply (auto simp add: of_nat_diff algebra_simps)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4314
    done
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4315
  have th0: "pochhammer (- (?a + ?b)) n = (- 1)^n * pochhammer (c - a) n"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59815
diff changeset
  4316
    unfolding pochhammer_minus
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4317
    by (simp add: algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4318
  have th1: "pochhammer (- ?b) n = (- 1)^n * pochhammer c n"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59815
diff changeset
  4319
    unfolding pochhammer_minus
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4320
    by simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4321
  have nz: "pochhammer c n \<noteq> 0" using c
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4322
    by (simp add: pochhammer_eq_0_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4323
  from Vandermonde_pochhammer_lemma[where a = "?a" and b="?b" and n=n, OF h, unfolded th0 th1]
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4324
  show ?thesis
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4325
    using nz by (simp add: field_simps sum_distrib_left)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4326
qed
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4327
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4328
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4329
subsubsection \<open>Formal trigonometric functions\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4330
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  4331
definition "fps_sin (c::'a::field_char_0) =
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4332
  Abs_fps (\<lambda>n. if even n then 0 else (- 1) ^((n - 1) div 2) * c^n /(of_nat (fact n)))"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4333
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  4334
definition "fps_cos (c::'a::field_char_0) =
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  4335
  Abs_fps (\<lambda>n. if even n then (- 1) ^ (n div 2) * c^n / (of_nat (fact n)) else 0)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4336
66466
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4337
lemma fps_sin_0 [simp]: "fps_sin 0 = 0"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4338
  by (intro fps_ext) (auto simp: fps_sin_def elim!: oddE)
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4339
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4340
lemma fps_cos_0 [simp]: "fps_cos 0 = 1"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4341
  by (intro fps_ext) (auto simp: fps_cos_def)
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4342
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  4343
lemma fps_sin_deriv:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4344
  "fps_deriv (fps_sin c) = fps_const c * fps_cos c"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4345
  (is "?lhs = ?rhs")
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  4346
proof (rule fps_ext)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4347
  fix n :: nat
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4348
  show "?lhs $ n = ?rhs $ n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4349
  proof (cases "even n")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4350
    case True
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4351
    have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4352
    also have "\<dots> = of_nat (n+1) * ((- 1)^(n div 2) * c^Suc n / of_nat (fact (Suc n)))"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4353
      using True by (simp add: fps_sin_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4354
    also have "\<dots> = (- 1)^(n div 2) * c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4355
      unfolding fact_Suc of_nat_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4356
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4357
    also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4358
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4359
    finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4360
      using True by (simp add: fps_cos_def field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4361
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4362
    case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4363
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4364
      by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4365
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4366
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4367
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4368
lemma fps_cos_deriv: "fps_deriv (fps_cos c) = fps_const (- c)* (fps_sin c)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4369
  (is "?lhs = ?rhs")
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  4370
proof (rule fps_ext)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4371
  have th0: "- ((- 1::'a) ^ n) = (- 1)^Suc n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4372
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4373
  show "?lhs $ n = ?rhs $ n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4374
  proof (cases "even n")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4375
    case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4376
    then have n0: "n \<noteq> 0" by presburger
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4377
    from False have th1: "Suc ((n - 1) div 2) = Suc n div 2"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4378
      by (cases n) simp_all
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4379
    have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4380
    also have "\<dots> = of_nat (n+1) * ((- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact (Suc n)))"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4381
      using False by (simp add: fps_cos_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4382
    also have "\<dots> = (- 1)^((n + 1) div 2)*c^Suc n * (of_nat (n+1) / (of_nat (Suc n) * of_nat (fact n)))"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4383
      unfolding fact_Suc of_nat_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4384
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4385
    also have "\<dots> = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4386
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4387
    also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4388
      unfolding th0 unfolding th1 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4389
    finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4390
      using False by (simp add: fps_sin_def field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4391
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4392
    case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4393
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4394
      by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4395
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4396
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4397
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4398
lemma fps_sin_cos_sum_of_squares: "(fps_cos c)\<^sup>2 + (fps_sin c)\<^sup>2 = 1"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4399
  (is "?lhs = _")
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  4400
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4401
  have "fps_deriv ?lhs = 0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4402
    apply (simp add:  fps_deriv_power fps_sin_deriv fps_cos_deriv)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4403
    apply (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4404
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4405
  then have "?lhs = fps_const (?lhs $ 0)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4406
    unfolding fps_deriv_eq_0_iff .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4407
  also have "\<dots> = 1"
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30952
diff changeset
  4408
    by (auto simp add: fps_eq_iff numeral_2_eq_2 fps_mult_nth fps_cos_def fps_sin_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4409
  finally show ?thesis .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4410
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4411
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4412
lemma fps_sin_nth_0 [simp]: "fps_sin c $ 0 = 0"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4413
  unfolding fps_sin_def by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4414
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4415
lemma fps_sin_nth_1 [simp]: "fps_sin c $ 1 = c"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4416
  unfolding fps_sin_def by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4417
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4418
lemma fps_sin_nth_add_2:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4419
    "fps_sin c $ (n + 2) = - (c * c * fps_sin c $ n / (of_nat (n + 1) * of_nat (n + 2)))"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4420
  unfolding fps_sin_def
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4421
  apply (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4422
  apply simp
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4423
  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4424
  apply simp
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4425
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4426
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4427
lemma fps_cos_nth_0 [simp]: "fps_cos c $ 0 = 1"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4428
  unfolding fps_cos_def by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4429
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4430
lemma fps_cos_nth_1 [simp]: "fps_cos c $ 1 = 0"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4431
  unfolding fps_cos_def by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4432
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4433
lemma fps_cos_nth_add_2:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4434
  "fps_cos c $ (n + 2) = - (c * c * fps_cos c $ n / (of_nat (n + 1) * of_nat (n + 2)))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4435
  unfolding fps_cos_def
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4436
  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4437
  apply simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4438
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4439
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4440
lemma nat_induct2: "P 0 \<Longrightarrow> P 1 \<Longrightarrow> (\<And>n. P n \<Longrightarrow> P (n + 2)) \<Longrightarrow> P (n::nat)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4441
  unfolding One_nat_def numeral_2_eq_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4442
  apply (induct n rule: nat_less_induct)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4443
  apply (case_tac n)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4444
  apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4445
  apply (rename_tac m)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4446
  apply (case_tac m)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4447
  apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4448
  apply (rename_tac k)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4449
  apply (case_tac k)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4450
  apply simp_all
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4451
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4452
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4453
lemma nat_add_1_add_1: "(n::nat) + 1 + 1 = n + 2"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4454
  by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4455
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4456
lemma eq_fps_sin:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4457
  assumes 0: "a $ 0 = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4458
    and 1: "a $ 1 = c"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4459
    and 2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)"
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4460
  shows "a = fps_sin c"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4461
  apply (rule fps_ext)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4462
  apply (induct_tac n rule: nat_induct2)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4463
  apply (simp add: 0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4464
  apply (simp add: 1 del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4465
  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4466
  apply (simp add: nat_add_1_add_1 fps_sin_nth_add_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4467
              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4468
  apply (subst minus_divide_left)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4469
  apply (subst nonzero_eq_divide_eq)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4470
  apply (simp del: of_nat_add of_nat_Suc)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4471
  apply (simp only: ac_simps)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4472
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4473
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4474
lemma eq_fps_cos:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4475
  assumes 0: "a $ 0 = 1"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4476
    and 1: "a $ 1 = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4477
    and 2: "fps_deriv (fps_deriv a) = - (fps_const c * fps_const c * a)"
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4478
  shows "a = fps_cos c"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4479
  apply (rule fps_ext)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4480
  apply (induct_tac n rule: nat_induct2)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4481
  apply (simp add: 0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4482
  apply (simp add: 1 del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4483
  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4484
  apply (simp add: nat_add_1_add_1 fps_cos_nth_add_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4485
              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4486
  apply (subst minus_divide_left)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4487
  apply (subst nonzero_eq_divide_eq)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4488
  apply (simp del: of_nat_add of_nat_Suc)
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4489
  apply (simp only: ac_simps)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4490
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4491
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4492
lemma mult_nth_0 [simp]: "(a * b) $ 0 = a $ 0 * b $ 0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4493
  by (simp add: fps_mult_nth)
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4494
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4495
lemma mult_nth_1 [simp]: "(a * b) $ 1 = a $ 0 * b $ 1 + a $ 1 * b $ 0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4496
  by (simp add: fps_mult_nth)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4497
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4498
lemma fps_sin_add: "fps_sin (a + b) = fps_sin a * fps_cos b + fps_cos a * fps_sin b"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4499
  apply (rule eq_fps_sin [symmetric], simp, simp del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4500
  apply (simp del: fps_const_neg fps_const_add fps_const_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4501
              add: fps_const_add [symmetric] fps_const_neg [symmetric]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4502
                   fps_sin_deriv fps_cos_deriv algebra_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4503
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4504
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4505
lemma fps_cos_add: "fps_cos (a + b) = fps_cos a * fps_cos b - fps_sin a * fps_sin b"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4506
  apply (rule eq_fps_cos [symmetric], simp, simp del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4507
  apply (simp del: fps_const_neg fps_const_add fps_const_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4508
              add: fps_const_add [symmetric] fps_const_neg [symmetric]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4509
                   fps_sin_deriv fps_cos_deriv algebra_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4510
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4511
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  4512
lemma fps_sin_even: "fps_sin (- c) = - fps_sin c"
56479
91958d4b30f7 revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
hoelzl
parents: 56410
diff changeset
  4513
  by (auto simp add: fps_eq_iff fps_sin_def)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  4514
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  4515
lemma fps_cos_odd: "fps_cos (- c) = fps_cos c"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  4516
  by (auto simp add: fps_eq_iff fps_cos_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  4517
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4518
definition "fps_tan c = fps_sin c / fps_cos c"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4519
66466
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4520
lemma fps_tan_0 [simp]: "fps_tan 0 = 0"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4521
  by (simp add: fps_tan_def)
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66373
diff changeset
  4522
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  4523
lemma fps_tan_deriv: "fps_deriv (fps_tan c) = fps_const c / (fps_cos c)\<^sup>2"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4524
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4525
  have th0: "fps_cos c $ 0 \<noteq> 0" by (simp add: fps_cos_def)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4526
  from this have "fps_cos c \<noteq> 0" by (intro notI) simp
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  4527
  hence "fps_deriv (fps_tan c) =
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4528
           fps_const c * (fps_cos c^2 + fps_sin c^2) / (fps_cos c^2)"
62102
877463945ce9 fix code generation for uniformity: uniformity is a non-computable pure data.
hoelzl
parents: 62101
diff changeset
  4529
    by (simp add: fps_tan_def fps_divide_deriv power2_eq_square algebra_simps
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4530
                  fps_sin_deriv fps_cos_deriv fps_const_neg[symmetric] div_mult_swap
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4531
             del: fps_const_neg)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4532
  also note fps_sin_cos_sum_of_squares
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4533
  finally show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4534
qed
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  4535
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4536
text \<open>Connection to @{const "fps_exp"} over the complex numbers --- Euler and de Moivre.\<close>
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4537
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4538
lemma fps_exp_ii_sin_cos: "fps_exp (\<i> * c) = fps_cos c + fps_const \<i> * fps_sin c"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4539
  (is "?l = ?r")
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4540
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4541
  have "?l $ n = ?r $ n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4542
  proof (cases "even n")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4543
    case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4544
    then obtain m where m: "n = 2 * m" ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4545
    show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4546
      by (simp add: m fps_sin_def fps_cos_def power_mult_distrib power_mult power_minus [of "c ^ 2"])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4547
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4548
    case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4549
    then obtain m where m: "n = 2 * m + 1" ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4550
    show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4551
      by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4552
        power_mult power_minus [of "c ^ 2"])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4553
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4554
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4555
    by (simp add: fps_eq_iff)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4556
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4557
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4558
lemma fps_exp_minus_ii_sin_cos: "fps_exp (- (\<i> * c)) = fps_cos c - fps_const \<i> * fps_sin c"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4559
  unfolding minus_mult_right fps_exp_ii_sin_cos by (simp add: fps_sin_even fps_cos_odd)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4560
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4561
lemma fps_const_minus: "fps_const (c::'a::group_add) - fps_const d = fps_const (c - d)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4562
  by (fact fps_const_sub)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4563
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4564
lemma fps_of_int: "fps_const (of_int c) = of_int c"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4565
  by (induction c) (simp_all add: fps_const_minus [symmetric] fps_of_nat fps_const_neg [symmetric] 
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4566
                             del: fps_const_minus fps_const_neg)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63040
diff changeset
  4567
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4568
lemma fps_deriv_of_int [simp]: "fps_deriv (of_int n) = 0"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4569
  by (simp add: fps_of_int [symmetric])
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4570
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  4571
lemma fps_numeral_fps_const: "numeral i = fps_const (numeral i :: 'a::comm_ring_1)"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4572
  by (fact numeral_fps_const) (* FIfps_XME: duplicate *)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4573
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4574
lemma fps_cos_fps_exp_ii: "fps_cos c = (fps_exp (\<i> * c) + fps_exp (- \<i> * c)) / fps_const 2"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4575
proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4576
  have th: "fps_cos c + fps_cos c = fps_cos c * fps_const 2"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  4577
    by (simp add: numeral_fps_const)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4578
  show ?thesis
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4579
    unfolding fps_exp_ii_sin_cos minus_mult_commute
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4580
    by (simp add: fps_sin_even fps_cos_odd numeral_fps_const fps_divide_unit fps_const_inverse th)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4581
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4582
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4583
lemma fps_sin_fps_exp_ii: "fps_sin c = (fps_exp (\<i> * c) - fps_exp (- \<i> * c)) / fps_const (2*\<i>)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4584
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63539
diff changeset
  4585
  have th: "fps_const \<i> * fps_sin c + fps_const \<i> * fps_sin c = fps_sin c * fps_const (2 * \<i>)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  4586
    by (simp add: fps_eq_iff numeral_fps_const)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4587
  show ?thesis
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4588
    unfolding fps_exp_ii_sin_cos minus_mult_commute
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4589
    by (simp add: fps_sin_even fps_cos_odd fps_divide_unit fps_const_inverse th)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4590
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4591
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4592
lemma fps_tan_fps_exp_ii:
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4593
  "fps_tan c = (fps_exp (\<i> * c) - fps_exp (- \<i> * c)) / 
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4594
      (fps_const \<i> * (fps_exp (\<i> * c) + fps_exp (- \<i> * c)))"
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4595
  unfolding fps_tan_def fps_sin_fps_exp_ii fps_cos_fps_exp_ii mult_minus_left fps_exp_neg
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4596
  apply (simp add: fps_divide_unit fps_inverse_mult fps_const_mult[symmetric] fps_const_inverse del: fps_const_mult)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4597
  apply simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4598
  done
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4599
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4600
lemma fps_demoivre:
63589
58aab4745e85 more symbols;
wenzelm
parents: 63539
diff changeset
  4601
  "(fps_cos a + fps_const \<i> * fps_sin a)^n =
58aab4745e85 more symbols;
wenzelm
parents: 63539
diff changeset
  4602
    fps_cos (of_nat n * a) + fps_const \<i> * fps_sin (of_nat n * a)"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4603
  unfolding fps_exp_ii_sin_cos[symmetric] fps_exp_power_mult
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4604
  by (simp add: ac_simps)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4605
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4606
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  4607
subsection \<open>Hypergeometric series\<close>
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4608
68442
nipkow
parents: 68072
diff changeset
  4609
definition "fps_hypergeo as bs (c::'a::field_char_0) =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4610
  Abs_fps (\<lambda>n. (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4611
    (foldl (\<lambda>r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4612
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4613
lemma fps_hypergeo_nth[simp]: "fps_hypergeo as bs c $ n =
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4614
  (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4615
    (foldl (\<lambda>r b. r * pochhammer b n) 1 bs * of_nat (fact n))"
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4616
  by (simp add: fps_hypergeo_def)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4617
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4618
lemma foldl_mult_start:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4619
  fixes v :: "'a::comm_ring_1"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4620
  shows "foldl (\<lambda>r x. r * f x) v as * x = foldl (\<lambda>r x. r * f x) (v * x) as "
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4621
  by (induct as arbitrary: x v) (auto simp add: algebra_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4622
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4623
lemma foldr_mult_foldl:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4624
  fixes v :: "'a::comm_ring_1"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4625
  shows "foldr (\<lambda>x r. r * f x) as v = foldl (\<lambda>r x. r * f x) v as"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4626
  by (induct as arbitrary: v) (auto simp add: foldl_mult_start)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4627
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4628
lemma fps_hypergeo_nth_alt:
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4629
  "fps_hypergeo as bs c $ n = foldr (\<lambda>a r. r * pochhammer a n) as (c ^ n) /
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4630
    foldr (\<lambda>b r. r * pochhammer b n) bs (of_nat (fact n))"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4631
  by (simp add: foldl_mult_start foldr_mult_foldl)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4632
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4633
lemma fps_hypergeo_fps_exp[simp]: "fps_hypergeo [] [] c = fps_exp c"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4634
  by (simp add: fps_eq_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4635
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4636
lemma fps_hypergeo_1_0[simp]: "fps_hypergeo [1] [] c = 1/(1 - fps_const c * fps_X)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4637
proof -
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4638
  let ?a = "(Abs_fps (\<lambda>n. 1)) oo (fps_const c * fps_X)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4639
  have th0: "(fps_const c * fps_X) $ 0 = 0" by simp
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4640
  show ?thesis unfolding gp[OF th0, symmetric]
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4641
    by (auto simp add: fps_eq_iff pochhammer_fact[symmetric]
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4642
      fps_compose_nth power_mult_distrib cond_value_iff sum.delta' cong del: if_weak_cong)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4643
qed
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4644
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4645
lemma fps_hypergeo_B[simp]: "fps_hypergeo [-a] [] (- 1) = fps_binomial a"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4646
  by (simp add: fps_eq_iff gbinomial_pochhammer algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4647
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4648
lemma fps_hypergeo_0[simp]: "fps_hypergeo as bs c $ 0 = 1"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4649
  apply simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4650
  apply (subgoal_tac "\<forall>as. foldl (\<lambda>(r::'a) (a::'a). r) 1 as = 1")
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4651
  apply auto
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4652
  apply (induct_tac as)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4653
  apply auto
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4654
  done
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4655
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4656
lemma foldl_prod_prod:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4657
  "foldl (\<lambda>(r::'b::comm_ring_1) (x::'a::comm_ring_1). r * f x) v as * foldl (\<lambda>r x. r * g x) w as =
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4658
    foldl (\<lambda>r x. r * f x * g x) (v * w) as"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4659
  by (induct as arbitrary: v w) (auto simp add: algebra_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4660
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4661
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4662
lemma fps_hypergeo_rec:
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4663
  "fps_hypergeo as bs c $ Suc n = ((foldl (\<lambda>r a. r* (a + of_nat n)) c as) /
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4664
    (foldl (\<lambda>r b. r * (b + of_nat n)) (of_nat (Suc n)) bs )) * fps_hypergeo as bs c $ n"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4665
  apply (simp del: of_nat_Suc of_nat_add fact_Suc)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4666
  apply (simp add: foldl_mult_start del: fact_Suc of_nat_Suc)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4667
  unfolding foldl_prod_prod[unfolded foldl_mult_start] pochhammer_Suc
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  4668
  apply (simp add: algebra_simps)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4669
  done
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4670
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4671
lemma fps_XD_nth[simp]: "fps_XD a $ n = (if n = 0 then 0 else of_nat n * a$n)"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4672
  by (simp add: fps_XD_def)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4673
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4674
lemma fps_XD_0th[simp]: "fps_XD a $ 0 = 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4675
  by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4676
lemma fps_XD_Suc[simp]:" fps_XD a $ Suc n = of_nat (Suc n) * a $ Suc n"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4677
  by simp
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4678
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4679
definition "fps_XDp c a = fps_XD a + fps_const c * a"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4680
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4681
lemma fps_XDp_nth[simp]: "fps_XDp c a $ n = (c + of_nat n) * a$n"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4682
  by (simp add: fps_XDp_def algebra_simps)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4683
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4684
lemma fps_XDp_commute: "fps_XDp b \<circ> fps_XDp (c::'a::comm_ring_1) = fps_XDp c \<circ> fps_XDp b"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4685
  by (auto simp add: fps_XDp_def fun_eq_iff fps_eq_iff algebra_simps)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4686
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4687
lemma fps_XDp0 [simp]: "fps_XDp 0 = fps_XD"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  4688
  by (simp add: fun_eq_iff fps_eq_iff)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4689
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4690
lemma fps_XDp_fps_integral [simp]: "fps_XDp 0 (fps_integral a c) = fps_X * a"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4691
  by (simp add: fps_eq_iff fps_integral_def)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4692
65396
b42167902f57 moved AFP material to Formal_Power_Series; renamed E/L/F in Formal_Power_Series
eberlm <eberlm@in.tum.de>
parents: 64786
diff changeset
  4693
lemma fps_hypergeo_minus_nat:
68442
nipkow
parents: 68072
diff changeset
  4694
  "fps_hypergeo [- of_nat n] [- of_nat (n + m)] (c::'a::field_char_0) $ k =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4695
    (if k \<le> n then
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4696
      pochhammer (- of_nat n) k * c ^ k / (pochhammer (- of_nat (n + m)) k * of_nat (fact k))
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4697
     else 0)"
68442
nipkow
parents: 68072
diff changeset
  4698
  "fps_hypergeo [- of_nat m] [- of_nat (m + n)] (c::'a::field_char_0) $ k =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4699
    (if k \<le> m then
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4700
      pochhammer (- of_nat m) k * c ^ k / (pochhammer (- of_nat (m + n)) k * of_nat (fact k))
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4701
     else 0)"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4702
  by (auto simp add: pochhammer_eq_0_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4703
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4704
lemma sum_eq_if: "sum f {(n::nat) .. m} = (if m < n then 0 else f n + sum f {n+1 .. m})"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4705
  apply simp
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4706
  apply (subst sum.insert[symmetric])
b9a1486e79be setsum -> sum
nipkow
parents: 64242
diff changeset
  4707
  apply (auto simp add: not_less sum_head_Suc)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4708
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4709
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4710
lemma pochhammer_rec_if: "pochhammer a n = (if n = 0 then 1 else a * pochhammer (a + 1) (n - 1))"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4711
  by (cases n) (simp_all add: pochhammer_rec)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4712
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4713
lemma fps_XDp_foldr_nth [simp]: "foldr (\<lambda>c r. fps_XDp c \<circ> r) cs (\<lambda>c. fps_XDp c a) c0 $ n =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4714
    foldr (\<lambda>c r. (c + of_nat n) * r) cs (c0 + of_nat n) * a$n"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4715
  by (induct cs arbitrary: c0) (auto simp add: algebra_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4716
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4717
lemma genric_fps_XDp_foldr_nth:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4718
  assumes f: "\<forall>n c a. f c a $ n = (of_nat n + k c) * a$n"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  4719
  shows "foldr (\<lambda>c r. f c \<circ> r) cs (\<lambda>c. g c a) c0 $ n =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4720
    foldr (\<lambda>c r. (k c + of_nat n) * r) cs (g c0 a $ n)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4721
  by (induct cs arbitrary: c0) (auto simp add: algebra_simps f)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4722
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4723
lemma dist_less_imp_nth_equal:
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4724
  assumes "dist f g < inverse (2 ^ i)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4725
    and"j \<le> i"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4726
  shows "f $ j = g $ j"
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  4727
proof (rule ccontr)
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  4728
  assume "f $ j \<noteq> g $ j"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4729
  hence "f \<noteq> g" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4730
  with assms have "i < subdegree (f - g)"
62390
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
  4731
    by (simp add: if_split_asm dist_fps_def)
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  4732
  also have "\<dots> \<le> j"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4733
    using \<open>f $ j \<noteq> g $ j\<close> by (intro subdegree_leI) simp_all
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  4734
  finally show False using \<open>j \<le> i\<close> by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4735
qed
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4736
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4737
lemma nth_equal_imp_dist_less:
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4738
  assumes "\<And>j. j \<le> i \<Longrightarrow> f $ j = g $ j"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4739
  shows "dist f g < inverse (2 ^ i)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4740
proof (cases "f = g")
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4741
  case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4742
  then show ?thesis by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4743
next
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4744
  case False
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4745
  with assms have "dist f g = inverse (2 ^ subdegree (f - g))"
62390
842917225d56 more canonical names
nipkow
parents: 62343
diff changeset
  4746
    by (simp add: if_split_asm dist_fps_def)
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4747
  moreover
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4748
  from assms and False have "i < subdegree (f - g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4749
    by (intro subdegree_greaterI) simp_all
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4750
  ultimately show ?thesis by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4751
qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4752
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4753
lemma dist_less_eq_nth_equal: "dist f g < inverse (2 ^ i) \<longleftrightarrow> (\<forall>j \<le> i. f $ j = g $ j)"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4754
  using dist_less_imp_nth_equal nth_equal_imp_dist_less by blast
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4755
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4756
instance fps :: (comm_ring_1) complete_space
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4757
proof
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4758
  fix fps_X :: "nat \<Rightarrow> 'a fps"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4759
  assume "Cauchy fps_X"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4760
  obtain M where M: "\<forall>i. \<forall>m \<ge> M i. \<forall>j \<le> i. fps_X (M i) $ j = fps_X m $ j"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4761
  proof -
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4762
    have "\<exists>M. \<forall>m \<ge> M. \<forall>j\<le>i. fps_X M $ j = fps_X m $ j" for i
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4763
    proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4764
      have "0 < inverse ((2::real)^i)" by simp
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4765
      from metric_CauchyD[OF \<open>Cauchy fps_X\<close> this] dist_less_imp_nth_equal
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4766
      show ?thesis by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4767
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4768
    then show ?thesis using that by metis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4769
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4770
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4771
  show "convergent fps_X"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4772
  proof (rule convergentI)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4773
    show "fps_X \<longlonglongrightarrow> Abs_fps (\<lambda>i. fps_X (M i) $ i)"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4774
      unfolding tendsto_iff
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4775
    proof safe
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4776
      fix e::real assume e: "0 < e"
61969
e01015e49041 more symbols;
wenzelm
parents: 61943
diff changeset
  4777
      have "(\<lambda>n. inverse (2 ^ n) :: real) \<longlonglongrightarrow> 0" by (rule LIMSEQ_inverse_realpow_zero) simp_all
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4778
      from this and e have "eventually (\<lambda>i. inverse (2 ^ i) < e) sequentially"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4779
        by (rule order_tendstoD)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4780
      then obtain i where "inverse (2 ^ i) < e"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4781
        by (auto simp: eventually_sequentially)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4782
      have "eventually (\<lambda>x. M i \<le> x) sequentially"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4783
        by (auto simp: eventually_sequentially)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4784
      then show "eventually (\<lambda>x. dist (fps_X x) (Abs_fps (\<lambda>i. fps_X (M i) $ i)) < e) sequentially"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4785
      proof eventually_elim
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4786
        fix x
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4787
        assume x: "M i \<le> x"
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4788
        have "fps_X (M i) $ j = fps_X (M j) $ j" if "j \<le> i" for j
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4789
          using M that by (metis nat_le_linear)
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4790
        with x have "dist (fps_X x) (Abs_fps (\<lambda>j. fps_X (M j) $ j)) < inverse (2 ^ i)"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4791
          using M by (force simp: dist_less_eq_nth_equal)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  4792
        also note \<open>inverse (2 ^ i) < e\<close>
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4793
        finally show "dist (fps_X x) (Abs_fps (\<lambda>j. fps_X (M j) $ j)) < e" .
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4794
      qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4795
    qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4796
  qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4797
qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4798
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4799
(* TODO: Figure out better notation for this thing *)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4800
no_notation fps_nth (infixl "$" 75)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4801
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4802
bundle fps_notation
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4803
begin
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4804
notation fps_nth (infixl "$" 75)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  4805
end
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4806
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
  4807
end