src/HOL/Library/Formal_Power_Series.thy
author haftmann
Sat, 19 Dec 2015 17:03:17 +0100
changeset 61891 76189756ff65
parent 61804 67381557cee8
child 61943 7fba644ed827
permissions -rw-r--r--
documentation on last state of the art concerning interpretation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 39302
diff changeset
     1
(*  Title:      HOL/Library/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
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
     8
imports Complex_Main "~~/src/HOL/Number_Theory/Euclidean_Algorithm"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
     9
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    10
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    11
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    12
subsection \<open>The type of formal power series\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    13
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 48757
diff changeset
    14
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
    15
  morphisms fps_nth Abs_fps
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    16
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    17
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    18
notation fps_nth (infixl "$" 75)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    19
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    20
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
    21
  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
    22
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    23
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
    24
  by (simp add: expand_fps_eq)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    25
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    26
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
    27
  by (simp add: Abs_fps_inverse)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    28
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    29
text \<open>Definition of the basic elements 0 and 1 and the basic operations of addition,
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    30
  negation and multiplication.\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    31
36409
d323e7773aa8 use new classes (linordered_)field_inverse_zero
haftmann
parents: 36350
diff changeset
    32
instantiation fps :: (zero) zero
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    33
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    34
  definition fps_zero_def: "0 = Abs_fps (\<lambda>n. 0)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    35
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    36
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    37
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    38
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
    39
  unfolding fps_zero_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    40
36409
d323e7773aa8 use new classes (linordered_)field_inverse_zero
haftmann
parents: 36350
diff changeset
    41
instantiation fps :: ("{one, zero}") one
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    42
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    43
  definition fps_one_def: "1 = Abs_fps (\<lambda>n. if n = 0 then 1 else 0)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    44
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    45
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    46
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
    47
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
    48
  unfolding fps_one_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    49
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
    50
instantiation fps :: (plus) plus
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    51
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    52
  definition fps_plus_def: "op + = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n + g $ n))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    53
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    54
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    55
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    56
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
    57
  unfolding fps_plus_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    58
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    59
instantiation fps :: (minus) minus
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    60
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    61
  definition fps_minus_def: "op - = (\<lambda>f g. Abs_fps (\<lambda>n. f $ n - g $ n))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    62
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    63
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    64
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    65
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
    66
  unfolding fps_minus_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    67
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    68
instantiation fps :: (uminus) uminus
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    69
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    70
  definition fps_uminus_def: "uminus = (\<lambda>f. Abs_fps (\<lambda>n. - (f $ n)))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    71
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    72
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    73
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    74
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
    75
  unfolding fps_uminus_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    76
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
    77
instantiation fps :: ("{comm_monoid_add, times}") times
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    78
begin
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    79
  definition fps_times_def: "op * = (\<lambda>f g. Abs_fps (\<lambda>n. \<Sum>i=0..n. f $ i * g $ (n - i)))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
    80
  instance ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    81
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    82
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
    83
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
    84
  unfolding fps_times_def by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    85
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
    86
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
    87
  unfolding fps_times_def by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
    88
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
    89
declare atLeastAtMost_iff [presburger]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
    90
declare Bex_def [presburger]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
    91
declare Ball_def [presburger]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
    92
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    93
lemma mult_delta_left:
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    94
  fixes x y :: "'a::mult_zero"
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    95
  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
    96
  by simp
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    97
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    98
lemma mult_delta_right:
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
    99
  fixes x y :: "'a::mult_zero"
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   100
  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
   101
  by simp
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
   102
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   103
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
   104
  by auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   105
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   106
lemma cond_application_beta: "(if b then f else g) x = (if b then f x else g x)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   107
  by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   108
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   109
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   110
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
   111
  they represent is a commutative ring with unity\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   112
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   113
instance fps :: (semigroup_add) semigroup_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   114
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   115
  fix a b c :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   116
  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
   117
    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
   118
qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   119
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   120
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
   121
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   122
  fix a b :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   123
  show "a + b = b + a"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   124
    by (simp add: fps_ext add.commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   125
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   126
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   127
lemma fps_mult_assoc_lemma:
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   128
  fixes k :: nat
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   129
    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
   130
  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
   131
         (\<Sum>j=0..k. \<Sum>i=0..k - j. f j i (n - j - i))"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   132
  by (induct k) (simp_all add: Suc_diff_le setsum.distrib add.assoc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   133
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   134
instance fps :: (semiring_0) semigroup_mult
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   135
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   136
  fix a b c :: "'a fps"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   137
  show "(a * b) * c = a * (b * c)"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   138
  proof (rule fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   139
    fix n :: nat
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   140
    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
   141
          (\<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
   142
      by (rule fps_mult_assoc_lemma)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   143
    then show "((a * b) * c) $ n = (a * (b * c)) $ n"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   144
      by (simp add: fps_mult_nth setsum_right_distrib setsum_left_distrib mult.assoc)
29911
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
qed
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   147
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   148
lemma fps_mult_commute_lemma:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   149
  fixes n :: nat
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
   150
    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
   151
  shows "(\<Sum>i=0..n. f i (n - i)) = (\<Sum>i=0..n. f (n - i) i)"
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
   152
  by (rule setsum.reindex_bij_witness[where i="op - n" and j="op - n"]) auto
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   153
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   154
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
   155
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   156
  fix a b :: "'a fps"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   157
  show "a * b = b * a"
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   158
  proof (rule fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   159
    fix n :: nat
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   160
    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
   161
      by (rule fps_mult_commute_lemma)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   162
    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
   163
      by (simp add: fps_mult_nth mult.commute)
29687
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
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   166
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   167
instance fps :: (monoid_add) monoid_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   168
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   169
  fix a :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   170
  show "0 + a = a" by (simp add: fps_ext)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   171
  show "a + 0 = a" by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   172
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   173
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   174
instance fps :: (comm_monoid_add) comm_monoid_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   175
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   176
  fix a :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   177
  show "0 + a = a" by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   178
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   179
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   180
instance fps :: (semiring_1) monoid_mult
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   181
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   182
  fix a :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   183
  show "1 * a = a"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   184
    by (simp add: fps_ext fps_mult_nth mult_delta_left setsum.delta)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   185
  show "a * 1 = a"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   186
    by (simp add: fps_ext fps_mult_nth mult_delta_right setsum.delta')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   187
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   188
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   189
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
   190
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   191
  fix a b c :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   192
  show "b = c" if "a + b = a + c"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   193
    using that by (simp add: expand_fps_eq)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   194
  show "b = c" if "b + a = c + a"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   195
    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
   196
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   197
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   198
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
   199
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   200
  fix a b c :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   201
  show "a + b - a = b"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   202
    by (simp add: expand_fps_eq)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   203
  show "a - b - c = a - (b + c)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   204
    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
   205
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   206
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   207
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
   208
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   209
instance fps :: (group_add) group_add
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   210
proof
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   211
  fix a b :: "'a fps"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   212
  show "- a + a = 0" by (simp add: fps_ext)
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
   213
  show "a + - b = a - b" by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   214
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   215
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   216
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
   217
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   218
  fix a b :: "'a fps"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   219
  show "- a + a = 0" by (simp add: fps_ext)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   220
  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
   221
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   222
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   223
instance fps :: (zero_neq_one) zero_neq_one
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60567
diff changeset
   224
  by standard (simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   225
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   226
instance fps :: (semiring_0) semiring
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   227
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   228
  fix a b c :: "'a fps"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   229
  show "(a + b) * c = a * c + b * c"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   230
    by (simp add: expand_fps_eq fps_mult_nth distrib_right setsum.distrib)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   231
  show "a * (b + c) = a * b + a * c"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   232
    by (simp add: expand_fps_eq fps_mult_nth distrib_left setsum.distrib)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   233
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   234
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   235
instance fps :: (semiring_0) semiring_0
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   236
proof
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   237
  fix a :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   238
  show "0 * a = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   239
    by (simp add: fps_ext fps_mult_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   240
  show "a * 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   241
    by (simp add: fps_ext fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   242
qed
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   243
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   244
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
   245
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   246
instance fps :: (semiring_1) semiring_1 ..
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   247
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   248
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   249
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
   250
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   251
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
   252
  by (simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   253
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   254
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
   255
  (is "?lhs \<longleftrightarrow> ?rhs")
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   256
proof
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   257
  let ?n = "LEAST n. f $ n \<noteq> 0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   258
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   259
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   260
    from that have "\<exists>n. f $ n \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   261
      by (simp add: fps_nonzero_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   262
    then have "f $ ?n \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   263
      by (rule LeastI_ex)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   264
    moreover have "\<forall>m<?n. f $ m = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   265
      by (auto dest: not_less_Least)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   266
    ultimately have "f $ ?n \<noteq> 0 \<and> (\<forall>m<?n. f $ m = 0)" ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   267
    then show ?thesis ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   268
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   269
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   270
    using that by (auto simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   271
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   272
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   273
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
   274
  by (rule expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   275
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   276
lemma fps_setsum_nth: "setsum f S $ n = setsum (\<lambda>k. (f k) $ n) S"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   277
proof (cases "finite S")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   278
  case True
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   279
  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
   280
next
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   281
  case False
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   282
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   283
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   284
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   285
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   286
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
   287
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   288
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
   289
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   290
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
   291
  unfolding fps_const_def by simp
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   292
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   293
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
   294
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   295
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   296
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
   297
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   298
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   299
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
   300
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   301
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   302
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
   303
  by (simp add: fps_ext)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   304
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   305
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
   306
  by (simp add: fps_ext)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   307
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   308
lemma fps_const_mult[simp]: "fps_const (c::'a::ring) * fps_const d = fps_const (c * d)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   309
  by (simp add: fps_eq_iff fps_mult_nth setsum.neutral)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   310
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   311
lemma fps_const_add_left: "fps_const (c::'a::monoid_add) + f =
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   312
    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
   313
  by (simp add: fps_ext)
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   314
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   315
lemma fps_const_add_right: "f + fps_const (c::'a::monoid_add) =
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   316
    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
   317
  by (simp add: fps_ext)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   318
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   319
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
   320
  unfolding fps_eq_iff fps_mult_nth
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   321
  by (simp add: fps_const_def mult_delta_left setsum.delta)
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   322
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   323
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
   324
  unfolding fps_eq_iff fps_mult_nth
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   325
  by (simp add: fps_const_def mult_delta_right setsum.delta')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   326
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   327
lemma fps_mult_left_const_nth [simp]: "(fps_const (c::'a::semiring_1) * f)$n = c* f$n"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   328
  by (simp add: fps_mult_nth mult_delta_left setsum.delta)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   329
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   330
lemma fps_mult_right_const_nth [simp]: "(f * fps_const (c::'a::semiring_1))$n = f$n * c"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   331
  by (simp add: fps_mult_nth mult_delta_right setsum.delta')
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   332
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   333
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   334
subsection \<open>Formal power series form an integral domain\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   335
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   336
instance fps :: (ring) ring ..
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_1) ring_1
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
   339
  by (intro_classes, auto simp add: distrib_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   340
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   341
instance fps :: (comm_ring_1) comm_ring_1
54230
b1d955791529 more simplification rules on unary and binary minus
haftmann
parents: 53374
diff changeset
   342
  by (intro_classes, auto simp add: distrib_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   343
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   344
instance fps :: (ring_no_zero_divisors) ring_no_zero_divisors
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   345
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   346
  fix a b :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   347
  assume "a \<noteq> 0" and "b \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   348
  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
   349
    unfolding fps_nonzero_nth_minimal
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   350
    by blast+
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   351
  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
   352
    by (rule fps_mult_nth)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   353
  also have "\<dots> = (a $ i * b $ (i + j - i)) + (\<Sum>k\<in>{0..i+j} - {i}. a $ k * b $ (i + j - k))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   354
    by (rule setsum.remove) simp_all
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   355
  also have "(\<Sum>k\<in>{0..i+j}-{i}. a $ k * b $ (i + j - k)) = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   356
  proof (rule setsum.neutral [rule_format])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   357
    fix k assume "k \<in> {0..i+j} - {i}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   358
    then have "k < i \<or> i+j-k < j"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   359
      by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   360
    then show "a $ k * b $ (i + j - k) = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   361
      using i j by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   362
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   363
  also have "a $ i * b $ (i + j - i) + 0 = a $ i * b $ j"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   364
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   365
  also have "a $ i * b $ j \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   366
    using i j by simp
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   367
  finally have "(a*b) $ (i+j) \<noteq> 0" .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   368
  then show "a * b \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   369
    unfolding fps_nonzero_nth by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   370
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   371
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   372
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
   373
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
   374
instance fps :: (idom) idom ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   375
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   376
lemma numeral_fps_const: "numeral k = fps_const (numeral k)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   377
  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
   378
    fps_const_add [symmetric])
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   379
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   380
lemma neg_numeral_fps_const:
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   381
  "(- numeral k :: 'a :: ring_1 fps) = fps_const (- numeral k)"
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   382
  by (simp add: numeral_fps_const)
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
   383
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   384
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
   385
  by (simp add: numeral_fps_const)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   386
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   387
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
   388
  by (simp add: numeral_fps_const)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   389
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   390
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   391
subsection \<open>The eXtractor series X\<close>
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   392
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   393
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
   394
  by (induct n) auto
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   395
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   396
definition "X = Abs_fps (\<lambda>n. if n = 1 then 1 else 0)"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   397
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   398
lemma X_mult_nth [simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   399
  "(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
   400
proof (cases "n = 0")
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   401
  case False
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   402
  have "(X * f) $n = (\<Sum>i = 0..n. X $ i * f $ (n - i))"
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   403
    by (simp add: fps_mult_nth)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   404
  also have "\<dots> = f $ (n - 1)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   405
    using False by (simp add: X_def mult_delta_left setsum.delta)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   406
  finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   407
    using False by simp
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   408
next
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
   409
  case True
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   410
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   411
    by (simp add: fps_mult_nth X_def)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   412
qed
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   413
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   414
lemma X_mult_right_nth[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   415
    "((f :: 'a::comm_semiring_1 fps) * X) $n = (if n = 0 then 0 else f $ (n - 1))"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   416
  by (metis X_mult_nth mult.commute)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   417
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   418
lemma X_power_iff: "X^k = Abs_fps (\<lambda>n. if n = k then 1::'a::comm_ring_1 else 0)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   419
proof (induct k)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   420
  case 0
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
   421
  then show ?case by (simp add: X_def fps_eq_iff)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   422
next
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   423
  case (Suc k)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   424
  have "(X^Suc k) $ m = (if m = Suc k then 1::'a else 0)" for m
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   425
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   426
    have "(X^Suc k) $ m = (if m = 0 then 0 else (X^k) $ (m - 1))"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   427
      by (simp del: One_nat_def)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   428
    then show ?thesis
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   429
      using Suc.hyps by (auto cong del: if_weak_cong)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   430
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   431
  then show ?case
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   432
    by (simp add: fps_eq_iff)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   433
qed
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   434
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   435
lemma X_nth[simp]: "X$n = (if n = 1 then 1 else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   436
  by (simp add: X_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   437
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   438
lemma X_power_nth[simp]: "(X^k) $n = (if n = k then 1 else 0::'a::comm_ring_1)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   439
  by (simp add: X_power_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   440
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   441
lemma X_power_mult_nth: "(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
   442
  apply (induct k arbitrary: n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   443
  apply simp
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   444
  unfolding power_Suc mult.assoc
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   445
  apply (case_tac n)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   446
  apply auto
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   447
  done
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   448
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   449
lemma X_power_mult_right_nth:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   450
    "((f :: 'a::comm_ring_1 fps) * X^k) $n = (if n < k then 0 else f $ (n - k))"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
   451
  by (metis X_power_mult_nth mult.commute)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   452
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   453
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   454
lemma X_neq_fps_const [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> fps_const c"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   455
proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   456
  assume "(X::'a fps) = fps_const (c::'a)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   457
  hence "X$1 = (fps_const (c::'a))$1" by (simp only:)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   458
  thus False by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   459
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   460
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   461
lemma X_neq_zero [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   462
  by (simp only: fps_const_0_eq_0[symmetric] X_neq_fps_const) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   463
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   464
lemma X_neq_one [simp]: "(X :: 'a :: zero_neq_one fps) \<noteq> 1"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   465
  by (simp only: fps_const_1_eq_1[symmetric] X_neq_fps_const) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   466
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   467
lemma X_neq_numeral [simp]: "(X :: 'a :: {semiring_1,zero_neq_one} fps) \<noteq> numeral c"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   468
  by (simp only: numeral_fps_const X_neq_fps_const) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   469
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   470
lemma X_pow_eq_X_pow_iff [simp]: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   471
  "(X :: ('a :: {comm_ring_1}) fps) ^ m = X ^ n \<longleftrightarrow> m = n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   472
proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   473
  assume "(X :: 'a fps) ^ m = X ^ n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   474
  hence "(X :: 'a fps) ^ m $ m = X ^ n $ m" by (simp only:)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   475
  thus "m = n" by (simp split: split_if_asm)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   476
qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   477
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   478
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
   479
subsection \<open>Subdegrees\<close>  
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   480
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   481
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
   482
  "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
   483
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   484
lemma subdegreeI:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   485
  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
   486
  shows   "subdegree f = d"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   487
proof-
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   488
  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
   489
  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
   490
  proof (rule Least_equality)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   491
    fix e assume "f $ e \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   492
    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
   493
    thus "e \<ge> d" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   494
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   495
  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
   496
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   497
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   498
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
   499
proof-
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   500
  assume "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   501
  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
   502
  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
   503
  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
   504
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   505
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   506
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   507
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
   508
proof (cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   509
  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
   510
  note less
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   511
  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
   512
  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
   513
qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   514
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   515
lemma subdegree_geI:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   516
  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
   517
  shows   "subdegree f \<ge> n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   518
proof (rule ccontr)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   519
  assume "\<not>(subdegree f \<ge> n)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   520
  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
   521
  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
   522
  ultimately show False by contradiction
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   523
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   524
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   525
lemma subdegree_greaterI:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   526
  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
   527
  shows   "subdegree f > n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   528
proof (rule ccontr)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   529
  assume "\<not>(subdegree f > n)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   530
  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
   531
  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
   532
  ultimately show False by contradiction
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   533
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   534
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   535
lemma subdegree_leI:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   536
  "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
   537
  by (rule leI) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   538
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   539
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   540
lemma subdegree_0 [simp]: "subdegree 0 = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   541
  by (simp add: subdegree_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   542
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   543
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
   544
  by (auto intro!: subdegreeI)
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
lemma subdegree_X [simp]: "subdegree (X :: ('a :: zero_neq_one) fps) = 1"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   547
  by (auto intro!: subdegreeI simp: X_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   548
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   549
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
   550
  by (cases "c = 0") (auto intro!: subdegreeI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   551
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   552
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
   553
  by (simp add: numeral_fps_const)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   554
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   555
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
   556
proof (cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   557
  assume "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   558
  thus ?thesis
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   559
    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
   560
qed simp_all
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 [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
   563
  by (simp add: subdegree_eq_0_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   564
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   565
lemma nth_subdegree_mult [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   566
  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
   567
  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
   568
proof-
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   569
  let ?n = "subdegree f + subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   570
  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
   571
    by (simp add: fps_mult_nth)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   572
  also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   573
  proof (intro setsum.cong)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   574
    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
   575
    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
   576
    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
   577
      by (elim disjE conjE) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   578
  qed auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   579
  also have "... = f $ subdegree f * g $ subdegree g" by (simp add: setsum.delta)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   580
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   581
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   582
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   583
lemma subdegree_mult [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   584
  assumes "f \<noteq> 0" "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   585
  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
   586
proof (rule subdegreeI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   587
  let ?n = "subdegree f + subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   588
  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
   589
  also have "... = (\<Sum>i=0..?n. if i = subdegree f then f$i * g$(?n-i) else 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   590
  proof (intro setsum.cong)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   591
    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
   592
    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
   593
    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
   594
      by (elim disjE conjE) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   595
  qed auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   596
  also have "... = f $ subdegree f * g $ subdegree g" by (simp add: setsum.delta)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   597
  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
   598
  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
   599
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   600
  fix m assume m: "m < subdegree f + subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   601
  have "(f * g) $ m = (\<Sum>i=0..m. f$i * g$(m-i))" by (simp add: fps_mult_nth) 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   602
  also have "... = (\<Sum>i=0..m. 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   603
  proof (rule setsum.cong)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   604
    fix i assume "i \<in> {0..m}"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   605
    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
   606
    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
   607
  qed auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   608
  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
   609
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   610
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   611
lemma subdegree_power [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   612
  "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
   613
  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
   614
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   615
lemma subdegree_uminus [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   616
  "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
   617
  by (simp add: subdegree_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   618
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   619
lemma subdegree_minus_commute [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   620
  "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
   621
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   622
  have "f - g = -(g - f)" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   623
  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
   624
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   625
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   626
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   627
lemma subdegree_add_ge:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   628
  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
   629
  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
   630
proof (rule subdegree_geI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   631
  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
   632
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   633
  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
   634
  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
   635
  thus "(f + g) $ i = 0" by force
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   636
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   637
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   638
lemma subdegree_add_eq1:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   639
  assumes "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   640
  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
   641
  shows   "subdegree (f + g) = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   642
proof (rule antisym[OF subdegree_leI])
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   643
  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
   644
    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
   645
  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
   646
  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
   647
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   648
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   649
lemma subdegree_add_eq2:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   650
  assumes "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   651
  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
   652
  shows   "subdegree (f + g) = subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   653
  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
   654
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   655
lemma subdegree_diff_eq1:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   656
  assumes "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   657
  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
   658
  shows   "subdegree (f - g) = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   659
  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
   660
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   661
lemma subdegree_diff_eq2:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   662
  assumes "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   663
  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
   664
  shows   "subdegree (f - g) = subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   665
  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
   666
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   667
lemma subdegree_diff_ge [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   668
  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
   669
  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
   670
  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
   671
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   672
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
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   675
subsection \<open>Shifting and slicing\<close>
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   676
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   677
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
   678
  "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
   679
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   680
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
   681
  by (simp add: fps_shift_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   682
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   683
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
   684
  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
   685
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   686
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
   687
  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
   688
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   689
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
   690
  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
   691
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   692
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
   693
  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
   694
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   695
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
   696
  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
   697
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   698
lemma fps_shift_X_power [simp]: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   699
  "n \<le> m \<Longrightarrow> fps_shift n (X ^ m) = (X ^ (m - n) ::'a::comm_ring_1 fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   700
  by (intro fps_ext) (auto simp: 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_times_X_power:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   703
  "n \<le> subdegree f \<Longrightarrow> fps_shift n f * X ^ n = (f :: 'a :: comm_ring_1 fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   704
  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   705
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   706
lemma fps_shift_times_X_power' [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   707
  "fps_shift n (f * X^n) = (f :: 'a :: comm_ring_1 fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   708
  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   709
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   710
lemma fps_shift_times_X_power'':
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   711
  "m \<le> n \<Longrightarrow> fps_shift n (f * X^m) = fps_shift (n - m) (f :: 'a :: comm_ring_1 fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   712
  by (intro fps_ext) (auto simp: X_power_mult_right_nth nth_less_subdegree_zero)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   713
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   714
lemma fps_shift_subdegree [simp]: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   715
  "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
   716
  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
   717
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   718
lemma subdegree_decompose:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   719
  "f = fps_shift (subdegree f) f * X ^ subdegree (f :: ('a :: comm_ring_1) fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   720
  by (rule fps_ext) (auto simp: X_power_mult_right_nth)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   721
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   722
lemma subdegree_decompose':
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   723
  "n \<le> subdegree (f :: ('a :: comm_ring_1) fps) \<Longrightarrow> f = fps_shift n f * X^n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   724
  by (rule fps_ext) (auto simp: X_power_mult_right_nth intro!: nth_less_subdegree_zero)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   725
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   726
lemma fps_shift_fps_shift:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   727
  "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
   728
  by (rule fps_ext) (simp add: add_ac)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   729
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   730
lemma fps_shift_add:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   731
  "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
   732
  by (simp add: fps_eq_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   733
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   734
lemma fps_shift_mult:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   735
  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
   736
  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
   737
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   738
  from assms have "g = fps_shift n g * X^n" by (rule subdegree_decompose')
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   739
  also have "h * ... = (h * fps_shift n g) * X^n" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   740
  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
   741
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   742
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   743
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   744
lemma fps_shift_mult_right:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   745
  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
   746
  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
   747
  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
   748
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   749
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
   750
  by (cases "f = 0") auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   751
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   752
lemma fps_shift_subdegree_zero_iff [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   753
  "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
   754
  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
   755
     (simp_all del: nth_subdegree_zero_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   756
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   757
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   758
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
   759
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   760
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
   761
  unfolding fps_cutoff_def by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   762
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   763
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
   764
proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   765
  assume A: "fps_cutoff n f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   766
  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
   767
  proof (cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   768
    assume "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   769
    with A have "n \<le> subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   770
      by (intro subdegree_geI) (auto simp: fps_eq_iff split: split_if_asm)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   771
    thus ?thesis ..
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   772
  qed simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   773
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
   774
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   775
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
   776
  by (simp add: fps_eq_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   777
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   778
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
   779
  by (simp add: fps_eq_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   780
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   781
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
   782
  by (simp add: fps_eq_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   783
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   784
lemma fps_cutoff_fps_const: "fps_cutoff n (fps_const c) = (if n = 0 then 0 else fps_const c)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   785
  by (simp add: fps_eq_iff)  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   786
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   787
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
   788
  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
   789
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   790
lemma fps_shift_cutoff: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   791
  "fps_shift n (f :: ('a :: comm_ring_1) fps) * X^n + fps_cutoff n f = f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   792
  by (simp add: fps_eq_iff X_power_mult_right_nth)
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
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   795
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
   796
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   797
definition (in dist) "ball x r = {y. dist y x < r}"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   798
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   799
instantiation fps :: (comm_ring_1) dist
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   800
begin
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   801
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   802
definition
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   803
  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
   804
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   805
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
   806
  by (simp add: dist_fps_def)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   807
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   808
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
   809
  by (simp add: dist_fps_def)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   810
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   811
instance ..
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   812
30746
d6915b738bd9 fps made instance of number_ring
chaieb
parents: 30488
diff changeset
   813
end
d6915b738bd9 fps made instance of number_ring
chaieb
parents: 30488
diff changeset
   814
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   815
instantiation fps :: (comm_ring_1) metric_space
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   816
begin
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   817
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   818
definition open_fps_def: "open (S :: 'a fps set) = (\<forall>a \<in> S. \<exists>r. r >0 \<and> ball a r \<subseteq> S)"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   819
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   820
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   821
instance
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   822
proof
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   823
  show "open S = (\<forall>x\<in>S. \<exists>e>0. \<forall>y. dist y x < e \<longrightarrow> y \<in> S)" for S :: "'a fps set"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   824
    by (auto simp add: open_fps_def ball_def subset_eq)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   825
  show th: "dist a b = 0 \<longleftrightarrow> a = b" for a b :: "'a fps"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   826
    by (simp add: dist_fps_def split: split_if_asm)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   827
  then have th'[simp]: "dist a a = 0" for a :: "'a fps" by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   828
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   829
  fix a b c :: "'a fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   830
  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
   831
  then show "dist a b \<le> dist a c + dist b c"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   832
  proof cases
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   833
    case 1
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   834
    then show ?thesis by (simp add: dist_fps_def)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   835
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   836
    case 2
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   837
    then show ?thesis
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   838
      by (cases "c = a") (simp_all add: th dist_fps_sym)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   839
  next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
   840
    case neq: 3
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   841
    have False if "dist a b > dist a c + dist b c"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   842
    proof -
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   843
      let ?n = "subdegree (a - b)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   844
      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
   845
      with that have "dist a b > dist a c" and "dist a b > dist b c" by simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   846
      with neq have "?n < subdegree (a - c)" and "?n < subdegree (b - c)"  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   847
        by (simp_all add: dist_fps_def field_simps)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   848
      hence "(a - c) $ ?n = 0" and "(b - c) $ ?n = 0" 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   849
        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
   850
      hence "(a - b) $ ?n = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   851
      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
   852
      ultimately show False by contradiction
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   853
    qed
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   854
    thus ?thesis by (auto simp add: not_le[symmetric])
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   855
  qed
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   856
qed
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   857
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   858
end
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   859
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   860
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   861
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
   862
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   863
lemma reals_power_lt_ex:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   864
  fixes x y :: real
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   865
  assumes xp: "x > 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   866
    and y1: "y > 1"
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   867
  shows "\<exists>k>0. (1/y)^k < x"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   868
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   869
  have yp: "y > 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   870
    using y1 by simp
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   871
  from reals_Archimedean2[of "max 0 (- log y x) + 1"]
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   872
  obtain k :: nat where k: "real k > max 0 (- log y x) + 1"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   873
    by blast
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   874
  from k have kp: "k > 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   875
    by simp
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   876
  from k have "real k > - log y x"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   877
    by simp
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   878
  then have "ln y * real k > - ln x"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   879
    unfolding log_def
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   880
    using ln_gt_zero_iff[OF yp] y1
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   881
    by (simp add: minus_divide_left field_simps del: minus_divide_left[symmetric])
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   882
  then have "ln y * real k + ln x > 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   883
    by simp
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   884
  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
   885
    by (simp add: ac_simps)
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   886
  then have "y ^ k * x > 1"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   887
    unfolding exp_zero exp_add exp_real_of_nat_mult exp_ln [OF xp] exp_ln [OF yp]
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   888
    by simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   889
  then have "x > (1 / y)^k" using yp
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
   890
    by (simp add: field_simps)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   891
  then show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   892
    using kp by blast
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   893
qed
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   894
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
   895
lemma fps_sum_rep_nth: "(setsum (\<lambda>i. fps_const(a$i)*X^i) {0..m})$n =
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   896
    (if n \<le> m then a$n else 0::'a::comm_ring_1)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   897
  apply (auto simp add: fps_setsum_nth cond_value_iff cong del: if_weak_cong)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   898
  apply (simp add: setsum.delta')
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
   899
  done
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   900
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
   901
lemma fps_notation: "(\<lambda>n. setsum (\<lambda>i. fps_const(a$i) * X^i) {0..n}) ----> a"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
   902
  (is "?s ----> a")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   903
proof -
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   904
  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
   905
  proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   906
    obtain n0 where n0: "(1/2)^n0 < r" "n0 > 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   907
      using reals_power_lt_ex[OF \<open>r > 0\<close>, of 2] by auto
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   908
    show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   909
    proof -
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   910
      have "dist (?s n) a < r" if nn0: "n \<ge> n0" for n
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   911
      proof -
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   912
        from that have thnn0: "(1/2)^n \<le> (1/2 :: real)^n0"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   913
          by (simp add: divide_simps)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   914
        show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   915
        proof (cases "?s n = a")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   916
          case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   917
          then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   918
            unfolding dist_eq_0_iff[of "?s n" a, symmetric]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   919
            using \<open>r > 0\<close> by (simp del: dist_eq_0_iff)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   920
        next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   921
          case False
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   922
          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
   923
            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
   924
          from False have kn: "subdegree (?s n - a) > n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   925
            by (intro subdegree_greaterI) (simp_all add: fps_sum_rep_nth)              
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   926
          then have "dist (?s n) a < (1/2)^n" 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   927
            by (simp add: field_simps dist_fps_def)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   928
          also have "\<dots> \<le> (1/2)^n0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   929
            using nn0 by (simp add: divide_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   930
          also have "\<dots> < r"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   931
            using n0 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   932
          finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   933
        qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   934
      qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   935
      then show ?thesis by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   936
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   937
  qed
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   938
  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
   939
    unfolding lim_sequentially by blast
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   940
qed
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
   941
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   942
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   943
subsection \<open>Inverses of formal power series\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   944
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
   945
declare setsum.cong[fundef_cong]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   946
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
   947
instantiation fps :: ("{comm_monoid_add,inverse,times,uminus}") inverse
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   948
begin
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   949
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   950
fun natfun_inverse:: "'a fps \<Rightarrow> nat \<Rightarrow> 'a"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   951
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   952
  "natfun_inverse f 0 = inverse (f$0)"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   953
| "natfun_inverse f n = - inverse (f$0) * setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   954
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   955
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
   956
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   957
definition fps_divide_def:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   958
  "f div g = (if g = 0 then 0 else 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
   959
     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
   960
     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
   961
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   962
instance ..
36311
ed3a87a7f977 epheremal replacement of field_simps by field_eq_simps; dropped old division_by_zero instance
haftmann
parents: 36309
diff changeset
   963
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   964
end
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   965
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   966
lemma fps_inverse_zero [simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   967
  "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
   968
  by (simp add: fps_ext fps_inverse_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   969
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   970
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
   971
  apply (auto simp add: expand_fps_eq fps_inverse_def)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   972
  apply (case_tac n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   973
  apply auto
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   974
  done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   975
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   976
lemma inverse_mult_eq_1 [intro]:
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   977
  assumes f0: "f$0 \<noteq> (0::'a::field)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   978
  shows "inverse f * f = 1"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   979
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   980
  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
   981
    by (simp add: mult.commute)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   982
  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
   983
    by (simp add: fps_inverse_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   984
  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
   985
    by (simp add: fps_mult_nth fps_inverse_def)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   986
  have "(inverse f * f)$n = 0" if np: "n > 0" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
   987
  proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   988
    from np have eq: "{0..n} = {0} \<union> {1 .. n}"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   989
      by auto
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   990
    have d: "{0} \<inter> {1 .. n} = {}"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
   991
      by auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   992
    from f0 np have th0: "- (inverse f $ n) =
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   993
      (setsum (\<lambda>i. f$i * natfun_inverse f (n - i)) {1..n}) / (f$0)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   994
      by (cases n) (simp_all add: divide_inverse fps_inverse_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
   995
    from th0[symmetric, unfolded nonzero_divide_eq_eq[OF f0]]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
   996
    have th1: "setsum (\<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
   997
      by (simp add: field_simps)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
   998
    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
   999
      unfolding fps_mult_nth ifn ..
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1000
    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
  1001
      by (simp add: eq)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1002
    also have "\<dots> = 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1003
      unfolding th1 ifn by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1004
    finally show ?thesis unfolding c .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1005
  qed
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1006
  with th0 show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1007
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1008
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1009
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1010
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
  1011
  by (simp add: fps_inverse_def nonzero_imp_inverse_nonzero)
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1012
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1013
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
  1014
  by (simp add: fps_inverse_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1015
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1016
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
  1017
proof
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1018
  assume A: "inverse f = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1019
  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
  1020
  thus "f $ 0 = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1021
qed (simp add: fps_inverse_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1022
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1023
lemma fps_inverse_idempotent[intro, simp]:
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1024
  assumes f0: "f$0 \<noteq> (0::'a::field)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1025
  shows "inverse (inverse f) = f"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1026
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1027
  from f0 have if0: "inverse f $ 0 \<noteq> 0" by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1028
  from inverse_mult_eq_1[OF f0] inverse_mult_eq_1[OF if0]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1029
  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
  1030
    by (simp add: ac_simps)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1031
  then show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1032
    using f0 unfolding mult_cancel_left by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1033
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1034
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1035
lemma fps_inverse_unique:
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1036
  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
  1037
  shows   "inverse f = g"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1038
proof -
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1039
  have f0: "f $ 0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1040
  proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1041
    assume "f $ 0 = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1042
    hence "0 = (f * g) $ 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1043
    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
  1044
    finally show False by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1045
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1046
  from inverse_mult_eq_1[OF this] fg
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1047
  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
  1048
    by (simp add: ac_simps)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1049
  then show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1050
    using f0
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1051
    unfolding mult_cancel_right
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  1052
    by (auto simp add: expand_fps_eq)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1053
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1054
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1055
lemma setsum_zero_lemma:
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1056
  fixes n::nat
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1057
  assumes "0 < n"
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1058
  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
  1059
proof -
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1060
  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
  1061
  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
  1062
  let ?h = "\<lambda>i. if i=n - 1 then - 1 else 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1063
  have th1: "setsum ?f {0..n} = setsum ?g {0..n}"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1064
    by (rule setsum.cong) auto
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1065
  have th2: "setsum ?g {0..n - 1} = setsum ?h {0..n - 1}"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1066
    apply (rule setsum.cong)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1067
    using assms
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1068
    apply auto
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1069
    done
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1070
  have eq: "{0 .. n} = {0.. n - 1} \<union> {n}"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1071
    by auto
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1072
  from assms have d: "{0.. n - 1} \<inter> {n} = {}"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1073
    by auto
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1074
  have f: "finite {0.. n - 1}" "finite {n}"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1075
    by auto
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1076
  show ?thesis
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1077
    unfolding th1
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1078
    apply (simp add: setsum.union_disjoint[OF f d, unfolded eq[symmetric]] del: One_nat_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1079
    unfolding th2
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1080
    apply (simp add: setsum.delta)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1081
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1082
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1083
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1084
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
  1085
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
  1086
  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
  1087
  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
  1088
  show ?thesis
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1089
  proof (rule fps_inverse_unique)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1090
    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
  1091
    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
  1092
    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
  1093
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1094
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1095
  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
  1096
  hence "inverse (f * g) = 0" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1097
  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
  1098
  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
  1099
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1100
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1101
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1102
lemma fps_inverse_gp: "inverse (Abs_fps(\<lambda>n. (1::'a::field))) =
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1103
    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
  1104
  apply (rule fps_inverse_unique)
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1105
  apply (simp_all add: fps_eq_iff fps_mult_nth setsum_zero_lemma)
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1106
  done
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1107
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1108
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
  1109
proof (cases "f$0 = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1110
  assume nz: "f$0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1111
  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
  1112
    by (subst subdegree_mult) auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1113
  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
  1114
  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
  1115
  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
  1116
qed (simp_all add: fps_inverse_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1117
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1118
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
  1119
proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1120
  assume "f dvd 1"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1121
  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
  1122
  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
  1123
  thus "f $ 0 \<noteq> 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1124
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1125
  assume A: "f $ 0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1126
  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
  1127
qed
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 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
  1130
  by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1131
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1132
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
  1133
  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
  1134
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1135
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1136
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1137
instantiation fps :: (field) ring_div
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1138
begin
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
definition fps_mod_def:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1141
  "f mod g = (if g = 0 then f else
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1142
     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
  1143
     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
  1144
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1145
lemma fps_mod_eq_zero: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1146
  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
  1147
  shows   "f mod g = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1148
  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
  1149
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1150
lemma fps_times_divide_eq: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1151
  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
  1152
  shows   "f div g * g = f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1153
proof (cases "f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1154
  assume nz: "f \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1155
  def n \<equiv> "subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1156
  def h \<equiv> "fps_shift n g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1157
  from assms have [simp]: "h $ 0 \<noteq> 0" unfolding h_def by (simp add: n_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1158
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1159
  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
  1160
    by (simp add: fps_divide_def Let_def h_def n_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1161
  also have "... = fps_shift n (f * inverse h) * X^n * h" unfolding h_def n_def
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1162
    by (subst subdegree_decompose[of g]) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1163
  also have "fps_shift n (f * inverse h) * X^n = f * inverse h"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1164
    by (rule fps_shift_times_X_power) (simp_all add: nz assms n_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1165
  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
  1166
  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
  1167
  finally show ?thesis by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1168
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
  1169
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1170
lemma 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1171
  assumes "g$0 \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1172
  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
  1173
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1174
  from assms have [simp]: "subdegree g = 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
  1175
  from assms show "f div g = f * inverse g" 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1176
    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
  1177
  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
  1178
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1179
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1180
context
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1181
begin
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1182
private lemma fps_divide_cancel_aux1:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1183
  assumes "h$0 \<noteq> (0 :: 'a :: field)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1184
  shows   "(h * f) div (h * g) = f div g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1185
proof (cases "g = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1186
  assume "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1187
  from assms have "h \<noteq> 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1188
  note nz [simp] = \<open>g \<noteq> 0\<close> \<open>h \<noteq> 0\<close>
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1189
  from assms have [simp]: "subdegree h = 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
  1190
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1191
  have "(h * f) div (h * g) = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1192
          fps_shift (subdegree g) (h * f * inverse (fps_shift (subdegree g) (h*g)))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1193
    by (simp add: fps_divide_def Let_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1194
  also have "h * f * inverse (fps_shift (subdegree g) (h*g)) = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1195
               (inverse h * h) * f * inverse (fps_shift (subdegree g) g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1196
    by (subst fps_shift_mult) (simp_all add: algebra_simps fps_inverse_mult)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1197
  also from assms have "inverse h * h = 1" by (rule inverse_mult_eq_1)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1198
  finally show "(h * f) div (h * g) = f div g" by (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
  1199
qed (simp_all add: fps_divide_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1200
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1201
private lemma fps_divide_cancel_aux2:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1202
  "(f * X^m) div (g * X^m) = f div (g :: 'a :: field fps)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1203
proof (cases "g = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1204
  assume [simp]: "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1205
  have "(f * X^m) div (g * X^m) = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1206
          fps_shift (subdegree g + m) (f*inverse (fps_shift (subdegree g + m) (g*X^m))*X^m)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1207
    by (simp add: fps_divide_def Let_def algebra_simps)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1208
  also have "... = f div g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1209
    by (simp add: fps_shift_times_X_power'' fps_divide_def Let_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1210
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1211
qed (simp_all add: fps_divide_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1212
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1213
instance proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1214
  fix f g :: "'a fps"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1215
  def n \<equiv> "subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1216
  def h \<equiv> "fps_shift n g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1217
  
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1218
  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
  1219
  proof (cases "g = 0 \<or> f = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1220
    assume "\<not>(g = 0 \<or> f = 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1221
    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
  1222
    show ?thesis
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1223
    proof (rule disjE[OF le_less_linear])
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1224
      assume "subdegree f \<ge> subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1225
      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
  1226
    next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1227
      assume "subdegree f < subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1228
      have g_decomp: "g = h * X^n" unfolding h_def n_def by (rule subdegree_decompose)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1229
      have "f div g * g + f mod g = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1230
              fps_shift n (f * inverse h) * g + fps_cutoff n (f * inverse h) * h" 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1231
        by (simp add: fps_mod_def fps_divide_def Let_def n_def h_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1232
      also have "... = h * (fps_shift n (f * inverse h) * X^n + fps_cutoff n (f * inverse h))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1233
        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
  1234
      also have "... = f * (inverse h * h)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1235
        by (subst fps_shift_cutoff) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1236
      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
  1237
      finally show ?thesis by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1238
    qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1239
  qed (auto simp: fps_mod_def fps_divide_def Let_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1240
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1241
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1242
  fix f g h :: "'a fps"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1243
  assume "h \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1244
  show "(h * f) div (h * g) = f div g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1245
  proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1246
    def m \<equiv> "subdegree h"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1247
    def h' \<equiv> "fps_shift m h"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1248
    have h_decomp: "h = h' * X ^ m" unfolding h'_def m_def by (rule subdegree_decompose)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1249
    from \<open>h \<noteq> 0\<close> have [simp]: "h'$0 \<noteq> 0" by (simp add: h'_def m_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1250
    have "(h * f) div (h * g) = (h' * f * X^m) div (h' * g * X^m)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1251
      by (simp add: h_decomp algebra_simps)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1252
    also have "... = f div g" by (simp add: fps_divide_cancel_aux1 fps_divide_cancel_aux2)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1253
    finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1254
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1255
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1256
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1257
  fix f g h :: "'a fps"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1258
  assume [simp]: "h \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1259
  def n \<equiv> "subdegree h"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1260
  def h' \<equiv> "fps_shift n h"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1261
  note dfs = n_def h'_def
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1262
  have "(f + g * h) div h = fps_shift n (f * inverse h') + fps_shift n (g * (h * inverse h'))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1263
    by (simp add: fps_divide_def Let_def dfs[symmetric] algebra_simps fps_shift_add)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1264
  also have "h * inverse h' = (inverse h' * h') * X^n"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1265
    by (subst subdegree_decompose) (simp_all add: dfs)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1266
  also have "... = X^n" by (subst inverse_mult_eq_1) (simp_all add: dfs)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1267
  also have "fps_shift n (g * X^n) = g" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1268
  also have "fps_shift n (f * inverse h') = f div h" 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1269
    by (simp add: fps_divide_def Let_def dfs)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1270
  finally show "(f + g * h) div h = g + f div h" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1271
qed (auto simp: fps_divide_def fps_mod_def Let_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1272
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1273
end
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1274
end
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1275
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1276
lemma subdegree_mod:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1277
  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
  1278
  shows   "subdegree (f mod g) = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1279
proof (cases "f div g * g = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1280
  assume "f div g * g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1281
  hence [simp]: "f div g \<noteq> 0" "g \<noteq> 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1282
  from mod_div_equality[of f g] have "f mod g = f - f div g * g" by (simp add: algebra_simps)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1283
  also from assms have "subdegree ... = subdegree f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1284
    by (intro subdegree_diff_eq1) simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1285
  finally show ?thesis .
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1286
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1287
  assume zero: "f div g * g = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1288
  from mod_div_equality[of f g] have "f mod g = f - f div g * g" by (simp add: algebra_simps)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1289
  also note zero
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1290
  finally show ?thesis by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1291
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1292
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1293
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
  1294
  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
  1295
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1296
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1297
lemma dvd_imp_subdegree_le: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1298
  "(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
  1299
  by (auto elim: dvdE)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1300
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1301
lemma fps_dvd_iff: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1302
  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
  1303
  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
  1304
proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1305
  assume "subdegree f \<le> subdegree g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1306
  with assms have "g mod f = 0" 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1307
    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
  1308
  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
  1309
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
  1310
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1311
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
  1312
  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
  1313
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1314
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
  1315
  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
  1316
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1317
lemma inverse_fps_numeral: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1318
  "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
  1319
  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
  1320
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1321
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1322
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1323
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1324
instantiation fps :: (field) normalization_semidom
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1325
begin
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1326
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1327
definition fps_unit_factor_def [simp]: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1328
  "unit_factor f = fps_shift (subdegree f) f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1329
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1330
definition fps_normalize_def [simp]:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1331
  "normalize f = (if f = 0 then 0 else X ^ subdegree f)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1332
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1333
instance proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1334
  fix f :: "'a fps"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1335
  show "unit_factor f * normalize f = f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1336
    by (simp add: fps_shift_times_X_power)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1337
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1338
  fix f g :: "'a fps"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1339
  show "unit_factor (f * g) = unit_factor f * unit_factor g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1340
  proof (cases "f = 0 \<or> g = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1341
    assume "\<not>(f = 0 \<or> g = 0)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1342
    thus "unit_factor (f * g) = unit_factor f * unit_factor g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1343
    unfolding fps_unit_factor_def
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1344
      by (auto simp: fps_shift_fps_shift fps_shift_mult fps_shift_mult_right)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1345
  qed auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1346
qed auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1347
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1348
end
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1349
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1350
instance fps :: (field) algebraic_semidom ..
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1351
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1352
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1353
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
  1354
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1355
instantiation fps :: (field) euclidean_ring
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1356
begin
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1357
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1358
definition fps_euclidean_size_def: 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1359
  "euclidean_size f = (if f = 0 then 0 else Suc (subdegree f))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1360
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1361
instance proof
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1362
  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
  1363
  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
  1364
    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
  1365
  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
  1366
    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
  1367
    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
  1368
    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
  1369
    done
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1370
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1371
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1372
end
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1373
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1374
instantiation fps :: (field) euclidean_ring_gcd
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1375
begin
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1376
definition fps_gcd_def: "(gcd :: 'a fps \<Rightarrow> _) = gcd_eucl"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1377
definition fps_lcm_def: "(lcm :: 'a fps \<Rightarrow> _) = lcm_eucl"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1378
definition fps_Gcd_def: "(Gcd :: 'a fps set \<Rightarrow> _) = Gcd_eucl"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1379
definition fps_Lcm_def: "(Lcm :: 'a fps set \<Rightarrow> _) = Lcm_eucl"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1380
instance by intro_classes (simp_all add: fps_gcd_def fps_lcm_def fps_Gcd_def fps_Lcm_def)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1381
end
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1382
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1383
lemma fps_gcd:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1384
  assumes [simp]: "f \<noteq> 0" "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1385
  shows   "gcd f g = X ^ min (subdegree f) (subdegree g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1386
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1387
  let ?m = "min (subdegree f) (subdegree g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1388
  show "gcd f g = X ^ ?m"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1389
  proof (rule sym, rule gcdI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1390
    fix d assume "d dvd f" "d dvd g"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1391
    thus "d dvd X ^ ?m" by (cases "d = 0") (auto simp: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1392
  qed (simp_all add: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1393
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1394
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1395
lemma fps_gcd_altdef: "gcd (f :: 'a :: field fps) g = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1396
  (if f = 0 \<and> g = 0 then 0 else
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1397
   if f = 0 then X ^ subdegree g else 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1398
   if g = 0 then X ^ subdegree f else 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1399
     X ^ min (subdegree f) (subdegree g))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1400
  by (simp add: fps_gcd)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1401
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1402
lemma fps_lcm:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1403
  assumes [simp]: "f \<noteq> 0" "g \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1404
  shows   "lcm f g = X ^ max (subdegree f) (subdegree g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1405
proof -
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1406
  let ?m = "max (subdegree f) (subdegree g)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1407
  show "lcm f g = X ^ ?m"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1408
  proof (rule sym, rule lcmI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1409
    fix d assume "f dvd d" "g dvd d"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1410
    thus "X ^ ?m dvd d" by (cases "d = 0") (auto simp: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1411
  qed (simp_all add: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1412
qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1413
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1414
lemma fps_lcm_altdef: "lcm (f :: 'a :: field fps) g = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1415
  (if f = 0 \<or> g = 0 then 0 else X ^ max (subdegree f) (subdegree g))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1416
  by (simp add: fps_lcm)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1417
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1418
lemma fps_Gcd:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1419
  assumes "A - {0} \<noteq> {}"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1420
  shows   "Gcd A = X ^ (INF f:A-{0}. subdegree f)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1421
proof (rule sym, rule GcdI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1422
  fix f assume "f \<in> A"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1423
  thus "X ^ (INF f:A - {0}. subdegree f) dvd f"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1424
    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
  1425
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1426
  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
  1427
  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
  1428
  with d[of f] have [simp]: "d \<noteq> 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1429
  from d assms have "subdegree d \<le> (INF f:A-{0}. subdegree f)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1430
    by (intro cINF_greatest) (auto simp: fps_dvd_iff[symmetric])
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1431
  with d assms show "d dvd X ^ (INF f:A-{0}. subdegree f)" by (simp add: fps_dvd_iff)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1432
qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1433
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1434
lemma fps_Gcd_altdef: "Gcd (A :: 'a :: field fps set) = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1435
  (if A \<subseteq> {0} then 0 else X ^ (INF f:A-{0}. subdegree f))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1436
  using fps_Gcd by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1437
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1438
lemma fps_Lcm:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1439
  assumes "A \<noteq> {}" "0 \<notin> A" "bdd_above (subdegree`A)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1440
  shows   "Lcm A = X ^ (SUP f:A. subdegree f)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1441
proof (rule sym, rule LcmI)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1442
  fix f assume "f \<in> A"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1443
  moreover from assms(3) have "bdd_above (subdegree ` A)" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1444
  ultimately show "f dvd X ^ (SUP f:A. subdegree f)" using assms(2)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1445
    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
  1446
next
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1447
  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
  1448
  from assms obtain f where f: "f \<in> A" "f \<noteq> 0" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1449
  show "X ^ (SUP f:A. subdegree f) dvd d"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1450
  proof (cases "d = 0")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1451
    assume "d \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1452
    moreover from d have "\<And>f. f \<in> A \<Longrightarrow> f \<noteq> 0 \<Longrightarrow> f dvd d" by blast
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1453
    ultimately have "subdegree d \<ge> (SUP f:A. subdegree f)" using assms
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1454
      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
  1455
    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
  1456
  qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1457
qed simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1458
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1459
lemma fps_Lcm_altdef:
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1460
  "Lcm (A :: 'a :: field fps set) = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1461
     (if 0 \<in> A \<or> \<not>bdd_above (subdegree`A) then 0 else
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1462
      if A = {} then 1 else X ^ (SUP f:A. subdegree f))"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1463
proof (cases "bdd_above (subdegree`A)")
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1464
  assume unbounded: "\<not>bdd_above (subdegree`A)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1465
  have "Lcm A = 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1466
  proof (rule ccontr)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1467
    assume "Lcm A \<noteq> 0"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1468
    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
  1469
      unfolding bdd_above_def by (auto simp: not_le)
61799
4cf66f21b764 isabelle update_cartouches -c -t;
wenzelm
parents: 61649
diff changeset
  1470
    moreover from this and \<open>Lcm A \<noteq> 0\<close> have "subdegree f \<le> subdegree (Lcm A)"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1471
      by (intro dvd_imp_subdegree_le) simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1472
    ultimately show False by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1473
  qed
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1474
  with unbounded show ?thesis by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1475
qed (simp_all add: fps_Lcm)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1476
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1477
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1478
subsection \<open>Formal Derivatives, and the MacLaurin theorem around 0\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1479
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1480
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
  1481
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1482
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
  1483
  by (simp add: fps_deriv_def)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1484
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1485
lemma fps_deriv_linear[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1486
  "fps_deriv (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1487
    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
  1488
  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
  1489
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1490
lemma fps_deriv_mult[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1491
  fixes f :: "'a::comm_ring_1 fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1492
  shows "fps_deriv (f * g) = f * fps_deriv g + fps_deriv f * g"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1493
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1494
  let ?D = "fps_deriv"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1495
  have "(f * ?D g + ?D f * g) $ n = ?D (f*g) $ n" for n
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1496
  proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1497
    let ?Zn = "{0 ..n}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1498
    let ?Zn1 = "{0 .. n + 1}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1499
    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
  1500
        of_nat (i+1)* f $ (i+1) * g $ (n - i)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1501
    let ?h = "\<lambda>i. of_nat i * g $ i * f $ ((n+1) - i) +
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1502
        of_nat i* f $ i * g $ ((n + 1) - i)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1503
    have s0: "setsum (\<lambda>i. of_nat i * f $ i * g $ (n + 1 - i)) ?Zn1 =
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1504
      setsum (\<lambda>i. of_nat (n + 1 - i) * f $ (n + 1 - i) * g $ i) ?Zn1"
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  1505
       by (rule setsum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1506
    have s1: "setsum (\<lambda>i. f $ i * g $ (n + 1 - i)) ?Zn1 =
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1507
      setsum (\<lambda>i. f $ (n + 1 - i) * g $ i) ?Zn1"
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  1508
       by (rule setsum.reindex_bij_witness[where i="op - (n + 1)" and j="op - (n + 1)"]) auto
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1509
    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
  1510
      by (simp only: mult.commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1511
    also have "\<dots> = (\<Sum>i = 0..n. ?g i)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1512
      by (simp add: fps_mult_nth setsum.distrib[symmetric])
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1513
    also have "\<dots> = setsum ?h {0..n+1}"
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  1514
      by (rule setsum.reindex_bij_witness_not_neutral
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  1515
            [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
  1516
    also have "\<dots> = (fps_deriv (f * g)) $ n"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1517
      apply (simp only: fps_deriv_nth fps_mult_nth setsum.distrib)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1518
      unfolding s0 s1
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1519
      unfolding setsum.distrib[symmetric] setsum_right_distrib
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1520
      apply (rule setsum.cong)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1521
      apply (auto simp add: of_nat_diff field_simps)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1522
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1523
    finally show ?thesis .
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1524
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1525
  then show ?thesis
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  1526
    unfolding fps_eq_iff by auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1527
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1528
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  1529
lemma fps_deriv_X[simp]: "fps_deriv X = 1"
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  1530
  by (simp add: fps_deriv_def X_def fps_eq_iff)
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  1531
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1532
lemma fps_deriv_neg[simp]:
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1533
  "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
  1534
  by (simp add: fps_eq_iff fps_deriv_def)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1535
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1536
lemma fps_deriv_add[simp]:
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1537
  "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
  1538
  using fps_deriv_linear[of 1 f 1 g] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1539
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1540
lemma fps_deriv_sub[simp]:
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1541
  "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
  1542
  using fps_deriv_add [of f "- g"] by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1543
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1544
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
  1545
  by (simp add: fps_ext fps_deriv_def fps_const_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1546
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1547
lemma fps_deriv_mult_const_left[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1548
  "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
  1549
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1550
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1551
lemma fps_deriv_0[simp]: "fps_deriv 0 = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1552
  by (simp add: fps_deriv_def fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1553
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1554
lemma fps_deriv_1[simp]: "fps_deriv 1 = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1555
  by (simp add: fps_deriv_def fps_eq_iff )
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1556
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1557
lemma fps_deriv_mult_const_right[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1558
  "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
  1559
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1560
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1561
lemma fps_deriv_setsum:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1562
  "fps_deriv (setsum f S) = setsum (\<lambda>i. fps_deriv (f i :: 'a::comm_ring_1 fps)) S"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1563
proof (cases "finite S")
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1564
  case False
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1565
  then show ?thesis by simp
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1566
next
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1567
  case True
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1568
  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1569
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1570
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1571
lemma fps_deriv_eq_0_iff [simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1572
  "fps_deriv f = 0 \<longleftrightarrow> f = fps_const (f$0 :: 'a::{idom,semiring_char_0})"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1573
  (is "?lhs \<longleftrightarrow> ?rhs")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1574
proof
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1575
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1576
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1577
    from that have "fps_deriv f = fps_deriv (fps_const (f$0))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1578
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1579
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1580
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1581
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1582
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1583
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1584
    from that have "\<forall>n. (fps_deriv f)$n = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1585
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1586
    then have "\<forall>n. f$(n+1) = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1587
      by (simp del: of_nat_Suc of_nat_add One_nat_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1588
    then show ?thesis
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1589
      apply (clarsimp simp add: fps_eq_iff fps_const_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1590
      apply (erule_tac x="n - 1" in allE)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1591
      apply simp
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1592
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1593
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1594
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1595
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1596
lemma fps_deriv_eq_iff:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1597
  fixes f :: "'a::{idom,semiring_char_0} fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1598
  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
  1599
proof -
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1600
  have "fps_deriv f = fps_deriv g \<longleftrightarrow> fps_deriv (f - g) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1601
    by simp
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1602
  also have "\<dots> \<longleftrightarrow> f - g = fps_const ((f - g) $ 0)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1603
    unfolding fps_deriv_eq_0_iff ..
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1604
  finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1605
    by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1606
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1607
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1608
lemma fps_deriv_eq_iff_ex:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1609
  "(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
  1610
  by (auto simp: fps_deriv_eq_iff)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1611
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1612
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1613
fun fps_nth_deriv :: "nat \<Rightarrow> 'a::semiring_1 fps \<Rightarrow> 'a fps"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1614
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1615
  "fps_nth_deriv 0 f = f"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1616
| "fps_nth_deriv (Suc n) f = fps_nth_deriv n (fps_deriv f)"
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_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
  1619
  by (induct n arbitrary: f) auto
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1620
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1621
lemma fps_nth_deriv_linear[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1622
  "fps_nth_deriv n (fps_const (a::'a::comm_semiring_1) * f + fps_const b * g) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1623
    fps_const a * fps_nth_deriv n f + fps_const b * fps_nth_deriv n g"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1624
  by (induct n arbitrary: f g) (auto simp add: fps_nth_deriv_commute)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1625
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1626
lemma fps_nth_deriv_neg[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1627
  "fps_nth_deriv n (- (f :: 'a::comm_ring_1 fps)) = - (fps_nth_deriv n f)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1628
  by (induct n arbitrary: f) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1629
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1630
lemma fps_nth_deriv_add[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1631
  "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
  1632
  using fps_nth_deriv_linear[of n 1 f 1 g] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1633
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1634
lemma fps_nth_deriv_sub[simp]:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1635
  "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
  1636
  using fps_nth_deriv_add [of n f "- g"] by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1637
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1638
lemma fps_nth_deriv_0[simp]: "fps_nth_deriv n 0 = 0"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1639
  by (induct n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1640
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1641
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
  1642
  by (induct n) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1643
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1644
lemma fps_nth_deriv_const[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1645
  "fps_nth_deriv n (fps_const c) = (if n = 0 then fps_const c else 0)"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1646
  by (cases n) simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1647
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1648
lemma fps_nth_deriv_mult_const_left[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1649
  "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
  1650
  using fps_nth_deriv_linear[of n "c" f 0 0 ] by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1651
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1652
lemma fps_nth_deriv_mult_const_right[simp]:
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1653
  "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
  1654
  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
  1655
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1656
lemma fps_nth_deriv_setsum:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1657
  "fps_nth_deriv n (setsum f S) = setsum (\<lambda>i. fps_nth_deriv n (f i :: 'a::comm_ring_1 fps)) S"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1658
proof (cases "finite S")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1659
  case True
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1660
  show ?thesis by (induct rule: finite_induct [OF True]) simp_all
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1661
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1662
  case False
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1663
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1664
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1665
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1666
lemma fps_deriv_maclauren_0:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1667
  "(fps_nth_deriv k (f :: 'a::comm_semiring_1 fps)) $ 0 = of_nat (fact k) * f $ k"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1668
  by (induct k arbitrary: f) (auto simp add: field_simps of_nat_mult)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1669
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1670
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1671
subsection \<open>Powers\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1672
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1673
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
  1674
  by (induct n) (auto simp add: expand_fps_eq fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1675
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1676
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
  1677
proof (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1678
  case 0
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1679
  then show ?case by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1680
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1681
  case (Suc n)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1682
  show ?case unfolding power_Suc fps_mult_nth
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1683
    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
  1684
    by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1685
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1686
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1687
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
  1688
  by (induct n) (auto simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1689
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1690
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
  1691
  by (induct n) (auto simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1692
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1693
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
  1694
  by (induct n) (auto simp add: fps_mult_nth)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1695
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1696
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
  1697
  apply (rule iffI)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1698
  apply (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1699
  apply (auto simp add: fps_mult_nth)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1700
  apply (rule startsby_zero_power, simp_all)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1701
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1702
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1703
lemma startsby_zero_power_prefix:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1704
  assumes a0: "a $ 0 = (0::'a::idom)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1705
  shows "\<forall>n < k. a ^ k $ n = 0"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1706
  using a0
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1707
proof (induct k rule: nat_less_induct)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1708
  fix k
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1709
  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
  1710
  show "\<forall>m<k. a ^ k $ m = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1711
  proof (cases k)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1712
    case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1713
    then show ?thesis by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1714
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1715
    case (Suc l)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1716
    have "a^k $ m = 0" if mk: "m < k" for m
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1717
    proof (cases "m = 0")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1718
      case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1719
      then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1720
        using startsby_zero_power[of a k] Suc a0 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1721
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1722
      case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1723
      have "a ^k $ m = (a^l * a) $m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1724
        by (simp add: Suc mult.commute)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1725
      also have "\<dots> = (\<Sum>i = 0..m. a ^ l $ i * a $ (m - i))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1726
        by (simp add: fps_mult_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1727
      also have "\<dots> = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1728
        apply (rule setsum.neutral)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1729
        apply auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1730
        apply (case_tac "x = m")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1731
        using a0 apply simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1732
        apply (rule H[rule_format])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1733
        using a0 Suc mk apply auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1734
        done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1735
      finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1736
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1737
    then show ?thesis by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1738
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1739
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1740
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1741
lemma startsby_zero_setsum_depends:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1742
  assumes a0: "a $0 = (0::'a::idom)"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1743
    and kn: "n \<ge> k"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1744
  shows "setsum (\<lambda>i. (a ^ i)$k) {0 .. n} = setsum (\<lambda>i. (a ^ i)$k) {0 .. k}"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1745
  apply (rule setsum.mono_neutral_right)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1746
  using kn
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1747
  apply auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1748
  apply (rule startsby_zero_power_prefix[rule_format, OF a0])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1749
  apply arith
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1750
  done
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1751
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1752
lemma startsby_zero_power_nth_same:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1753
  assumes a0: "a$0 = (0::'a::idom)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1754
  shows "a^n $ n = (a$1) ^ n"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1755
proof (induct n)
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1756
  case 0
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1757
  then show ?case by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1758
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1759
  case (Suc n)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1760
  have "a ^ Suc n $ (Suc n) = (a^n * a)$(Suc n)"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1761
    by (simp add: field_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1762
  also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {0.. Suc n}"
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1763
    by (simp add: fps_mult_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1764
  also have "\<dots> = setsum (\<lambda>i. a^n$i * a $ (Suc n - i)) {n .. Suc n}"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1765
    apply (rule setsum.mono_neutral_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1766
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1767
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1768
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1769
    apply (rule startsby_zero_power_prefix[rule_format, OF a0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1770
    apply arith
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1771
    done
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1772
  also have "\<dots> = a^n $ n * a$1"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1773
    using a0 by simp
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1774
  finally show ?case
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1775
    using Suc.hyps by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1776
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1777
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1778
lemma fps_inverse_power:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1779
  fixes a :: "'a::field fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1780
  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
  1781
  by (induction n) (simp_all add: fps_inverse_mult)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1782
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1783
lemma fps_deriv_power:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1784
  "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
  1785
  apply (induct n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1786
  apply (auto simp add: field_simps fps_const_add[symmetric] simp del: fps_const_add)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1787
  apply (case_tac n)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  1788
  apply (auto simp add: field_simps)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1789
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1790
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1791
lemma fps_inverse_deriv:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1792
  fixes a :: "'a::field fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1793
  assumes a0: "a$0 \<noteq> 0"
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1794
  shows "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1795
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1796
  from inverse_mult_eq_1[OF a0]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1797
  have "fps_deriv (inverse a * a) = 0" by simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  1798
  then have "inverse a * fps_deriv a + fps_deriv (inverse a) * a = 0"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  1799
    by simp
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  1800
  then have "inverse a * (inverse a * fps_deriv a + fps_deriv (inverse a) * a) = 0"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  1801
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1802
  with inverse_mult_eq_1[OF a0]
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1803
  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
  1804
    unfolding power2_eq_square
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1805
    apply (simp add: field_simps)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  1806
    apply (simp add: mult.assoc[symmetric])
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1807
    done
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1808
  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
  1809
      0 - fps_deriv a * (inverse a)\<^sup>2"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1810
    by simp
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1811
  then show "fps_deriv (inverse a) = - fps_deriv a * (inverse a)\<^sup>2"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1812
    by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1813
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1814
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1815
lemma fps_inverse_deriv':
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1816
  fixes a :: "'a::field fps"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1817
  assumes a0: "a $ 0 \<noteq> 0"
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  1818
  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
  1819
  using fps_inverse_deriv[OF a0] a0
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1820
  by (simp add: fps_divide_unit power2_eq_square fps_inverse_mult)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1821
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1822
lemma inverse_mult_eq_1':
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1823
  assumes f0: "f$0 \<noteq> (0::'a::field)"
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  1824
  shows "f * inverse f = 1"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  1825
  by (metis mult.commute inverse_mult_eq_1 f0)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1826
61804
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1827
(* FIXME: The last part of this proof should go through by simp once we have a proper
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1828
   theorem collection for simplifying division on rings *)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1829
lemma fps_divide_deriv:
61804
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1830
  assumes "b dvd (a :: 'a :: field fps)"
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1831
  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
  1832
proof -
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1833
  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
  1834
    by (drule sym) (simp add: mult.assoc)
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1835
  from assms have "a = a / b * b" by simp
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1836
  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
  1837
  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
  1838
    by (simp add: power2_eq_square algebra_simps)
67381557cee8 Generalised derivative rule for division on formal power series
eberlm
parents: 61799
diff changeset
  1839
  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
  1840
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1841
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1842
lemma fps_inverse_gp': "inverse (Abs_fps (\<lambda>n. 1::'a::field)) = 1 - X"
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  1843
  by (simp add: fps_inverse_gp fps_eq_iff X_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1844
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1845
lemma fps_nth_deriv_X[simp]: "fps_nth_deriv n X = (if n = 0 then X else if n=1 then 1 else 0)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1846
  by (cases n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1847
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1848
lemma fps_inverse_X_plus1: "inverse (1 + X) = Abs_fps (\<lambda>n. (- (1::'a::field)) ^ n)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1849
  (is "_ = ?r")
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1850
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1851
  have eq: "(1 + X) * ?r = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1852
    unfolding minus_one_power_iff
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1853
    by (auto simp add: field_simps fps_eq_iff)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1854
  show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1855
    by (auto simp add: eq intro: fps_inverse_unique)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1856
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1857
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1858
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1859
subsection \<open>Integration\<close>
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1860
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1861
definition fps_integral :: "'a::field_char_0 fps \<Rightarrow> 'a \<Rightarrow> 'a fps"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  1862
  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
  1863
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1864
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
  1865
  unfolding fps_integral_def fps_deriv_def
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1866
  by (simp add: fps_eq_iff del: of_nat_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1867
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1868
lemma fps_integral_linear:
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  1869
  "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
  1870
    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
  1871
  (is "?l = ?r")
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1872
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1873
  have "fps_deriv ?l = fps_deriv ?r"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1874
    by (simp add: fps_deriv_fps_integral)
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1875
  moreover have "?l$0 = ?r$0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1876
    by (simp add: fps_integral_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1877
  ultimately show ?thesis
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1878
    unfolding fps_deriv_eq_iff by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1879
qed
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1880
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1881
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1882
subsection \<open>Composition of FPSs\<close>
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1883
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1884
definition fps_compose :: "'a::semiring_1 fps \<Rightarrow> 'a fps \<Rightarrow> 'a fps"  (infixl "oo" 55)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1885
  where "a oo b = Abs_fps (\<lambda>n. setsum (\<lambda>i. a$i * (b^i$n)) {0..n})"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1886
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1887
lemma fps_compose_nth: "(a oo b)$n = setsum (\<lambda>i. a$i * (b^i$n)) {0..n}"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1888
  by (simp add: fps_compose_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1889
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1890
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
  1891
  by (simp add: fps_compose_nth)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  1892
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1893
lemma fps_compose_X[simp]: "a oo X = (a :: 'a::comm_ring_1 fps)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1894
  by (simp add: fps_ext fps_compose_def mult_delta_right setsum.delta')
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1895
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1896
lemma fps_const_compose[simp]: "fps_const (a::'a::comm_ring_1) oo b = fps_const a"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1897
  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1898
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1899
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
  1900
  unfolding numeral_fps_const by simp
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  1901
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1902
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
  1903
  unfolding neg_numeral_fps_const by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  1904
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1905
lemma X_fps_compose_startby0[simp]: "a$0 = 0 \<Longrightarrow> X oo a = (a :: 'a::comm_ring_1 fps)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  1906
  by (simp add: fps_eq_iff fps_compose_def mult_delta_left setsum.delta not_le)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1907
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1908
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1909
subsection \<open>Rules from Herbert Wilf's Generatingfunctionology\<close>
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1910
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1911
subsubsection \<open>Rule 1\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1912
  (* {a_{n+k}}_0^infty Corresponds to (f - setsum (\<lambda>i. a_i * x^i))/x^h, for h>0*)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1913
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1914
lemma fps_power_mult_eq_shift:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1915
  "X^Suc k * Abs_fps (\<lambda>n. a (n + Suc k)) =
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1916
    Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a::comm_ring_1) * X^i) {0 .. k}"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1917
  (is "?lhs = ?rhs")
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1918
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1919
  have "?lhs $ n = ?rhs $ n" for n :: nat
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1920
  proof -
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  1921
    have "?lhs $ n = (if n < Suc k then 0 else a n)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1922
      unfolding X_power_mult_nth by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1923
    also have "\<dots> = ?rhs $ n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1924
    proof (induct k)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1925
      case 0
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1926
      then show ?case
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1927
        by (simp add: fps_setsum_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1928
    next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1929
      case (Suc k)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1930
      have "(Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. Suc k})$n =
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  1931
        (Abs_fps a - setsum (\<lambda>i. fps_const (a i :: 'a) * X^i) {0 .. k} -
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  1932
          fps_const (a (Suc k)) * X^ Suc k) $ n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1933
        by (simp add: field_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1934
      also have "\<dots> = (if n < Suc k then 0 else a n) - (fps_const (a (Suc k)) * X^ Suc k)$n"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1935
        using Suc.hyps[symmetric] unfolding fps_sub_nth by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1936
      also have "\<dots> = (if n < Suc (Suc k) then 0 else a n)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1937
        unfolding X_power_mult_right_nth
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  1938
        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
  1939
        apply (rule cong[of a a, OF refl])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1940
        apply arith
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1941
        done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1942
      finally show ?case
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1943
        by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1944
    qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1945
    finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1946
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1947
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1948
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1949
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1950
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1951
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1952
subsubsection \<open>Rule 2\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1953
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1954
  (* 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
  1955
  (* If f reprents {a_n} and P is a polynomial, then
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1956
        P(xD) f represents {P(n) a_n}*)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1957
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1958
definition "XD = op * X \<circ> fps_deriv"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1959
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1960
lemma XD_add[simp]:"XD (a + b) = XD a + XD (b :: 'a::comm_ring_1 fps)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1961
  by (simp add: XD_def field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1962
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1963
lemma XD_mult_const[simp]:"XD (fps_const (c::'a::comm_ring_1) * a) = fps_const c * XD a"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  1964
  by (simp add: XD_def field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1965
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1966
lemma XD_linear[simp]: "XD (fps_const c * a + fps_const d * b) =
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1967
    fps_const c * XD a + fps_const d * XD (b :: 'a::comm_ring_1 fps)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1968
  by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1969
30952
7ab2716dd93b power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents: 30837
diff changeset
  1970
lemma XDN_linear:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1971
  "(XD ^^ n) (fps_const c * a + fps_const d * b) =
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1972
    fps_const c * (XD ^^ n) a + fps_const d * (XD ^^ n) (b :: 'a::comm_ring_1 fps)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  1973
  by (induct n) simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1974
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1975
lemma fps_mult_X_deriv_shift: "X* fps_deriv a = Abs_fps (\<lambda>n. of_nat n* a$n)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1976
  by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1977
30952
7ab2716dd93b power operation on functions with syntax o^; power operation on relations with syntax ^^
haftmann
parents: 30837
diff changeset
  1978
lemma fps_mult_XD_shift:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1979
  "(XD ^^ k) (a :: 'a::comm_ring_1 fps) = Abs_fps (\<lambda>n. (of_nat n ^ k) * a$n)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1980
  by (induct k arbitrary: a) (simp_all add: XD_def fps_eq_iff field_simps del: One_nat_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1981
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1982
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1983
subsubsection \<open>Rule 3\<close>
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1984
61585
a9599d3d7610 isabelle update_cartouches -c -t;
wenzelm
parents: 61552
diff changeset
  1985
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
  1986
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1987
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1988
subsubsection \<open>Rule 5 --- summation and "division" by (1 - X)\<close>
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
lemma fps_divide_X_minus1_setsum_lemma:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  1991
  "a = ((1::'a::comm_ring_1 fps) - X) * Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  1992
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  1993
  let ?sa = "Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1994
  have th0: "\<And>i. (1 - (X::'a fps)) $ i = (if i = 0 then 1 else if i = 1 then - 1 else 0)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  1995
    by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1996
  have "a$n = ((1 - X) * ?sa) $ n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1997
  proof (cases "n = 0")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1998
    case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  1999
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2000
      by (simp add: fps_mult_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2001
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2002
    case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2003
    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
  2004
      "{0..n - 1} \<union> {n} = {0..n}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2005
      by (auto simp: set_eq_iff)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2006
    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
  2007
      using False by simp_all
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2008
    have f: "finite {0}" "finite {1}" "finite {2 .. n}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2009
      "finite {0 .. n - 1}" "finite {n}" by simp_all
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2010
    have "((1 - X) * ?sa) $ n = setsum (\<lambda>i. (1 - X)$ i * ?sa $ (n - i)) {0 .. n}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2011
      by (simp add: fps_mult_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2012
    also have "\<dots> = a$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2013
      unfolding th0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2014
      unfolding setsum.union_disjoint[OF f(1) finite_UnI[OF f(2,3)] d(1), unfolded u(1)]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2015
      unfolding setsum.union_disjoint[OF f(2) f(3) d(2)]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2016
      apply (simp)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2017
      unfolding setsum.union_disjoint[OF f(4,5) d(3), unfolded u(3)]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2018
      apply simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2019
      done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2020
    finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2021
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2022
  qed
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2023
  then show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2024
    unfolding fps_eq_iff by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2025
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2026
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2027
lemma fps_divide_X_minus1_setsum:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2028
  "a /((1::'a::field fps) - X) = Abs_fps (\<lambda>n. setsum (\<lambda>i. a $ i) {0..n})"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2029
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2030
  let ?X = "1 - (X::'a fps)"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2031
  have th0: "?X $ 0 \<noteq> 0"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2032
    by simp
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2033
  have "a /?X = ?X *  Abs_fps (\<lambda>n::nat. setsum (op $ a) {0..n}) * inverse ?X"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2034
    using fps_divide_X_minus1_setsum_lemma[of a, symmetric] th0
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  2035
    by (simp add: fps_divide_def mult.assoc)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2036
  also have "\<dots> = (inverse ?X * ?X) * Abs_fps (\<lambda>n::nat. setsum (op $ a) {0..n}) "
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  2037
    by (simp add: ac_simps)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2038
  finally show ?thesis
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2039
    by (simp add: inverse_mult_eq_1[OF th0])
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2040
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2041
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2042
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2043
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
  2044
  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
  2045
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2046
definition "natpermute n k = {l :: nat list. length l = k \<and> listsum l = n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2047
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2048
lemma natlist_trivial_1: "natpermute n 1 = {[n]}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2049
  apply (auto simp add: natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2050
  apply (case_tac x)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2051
  apply auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2052
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2053
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2054
lemma append_natpermute_less_eq:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2055
  assumes "xs @ ys \<in> natpermute n k"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2056
  shows "listsum xs \<le> n"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2057
    and "listsum ys \<le> n"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2058
proof -
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2059
  from assms have "listsum (xs @ ys) = n"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2060
    by (simp add: natpermute_def)
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2061
  then have "listsum xs + listsum ys = n"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2062
    by simp
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2063
  then show "listsum xs \<le> n" and "listsum ys \<le> n"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2064
    by simp_all
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2065
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2066
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2067
lemma natpermute_split:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2068
  assumes "h \<le> k"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2069
  shows "natpermute n k =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2070
    (\<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
  2071
  (is "?L = ?R" is "_ = (\<Union>m \<in>{0..n}. ?S m)")
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2072
proof
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2073
  show "?R \<subseteq> ?L"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2074
  proof
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2075
    fix l
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2076
    assume l: "l \<in> ?R"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2077
    from l obtain m xs ys where h: "m \<in> {0..n}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2078
      and xs: "xs \<in> natpermute m h"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2079
      and ys: "ys \<in> natpermute (n - m) (k - h)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2080
      and leq: "l = xs@ys" by blast
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2081
    from xs have xs': "listsum xs = m"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2082
      by (simp add: natpermute_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2083
    from ys have ys': "listsum ys = n - m"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2084
      by (simp add: natpermute_def)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2085
    show "l \<in> ?L" using leq xs ys h
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2086
      apply (clarsimp simp add: natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2087
      unfolding xs' ys'
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2088
      using assms xs ys
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2089
      unfolding natpermute_def
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2090
      apply simp
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2091
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2092
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2093
  show "?L \<subseteq> ?R"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2094
  proof
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2095
    fix l
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2096
    assume l: "l \<in> natpermute n k"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2097
    let ?xs = "take h l"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2098
    let ?ys = "drop h l"
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2099
    let ?m = "listsum ?xs"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2100
    from l have ls: "listsum (?xs @ ?ys) = n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2101
      by (simp add: natpermute_def)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2102
    have xs: "?xs \<in> natpermute ?m h" using l assms
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2103
      by (simp add: natpermute_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2104
    have l_take_drop: "listsum l = listsum (take h l @ drop h l)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2105
      by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2106
    then have ys: "?ys \<in> natpermute (n - ?m) (k - h)"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2107
      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
  2108
    from ls have m: "?m \<in> {0..n}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2109
      by (simp add: l_take_drop del: append_take_drop_id)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2110
    from xs ys ls show "l \<in> ?R"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2111
      apply auto
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2112
      apply (rule bexI [where x = "?m"])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2113
      apply (rule exI [where x = "?xs"])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2114
      apply (rule exI [where x = "?ys"])
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2115
      using ls l
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2116
      apply (auto simp add: natpermute_def l_take_drop simp del: append_take_drop_id)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2117
      apply simp
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2118
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2119
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2120
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2121
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2122
lemma natpermute_0: "natpermute n 0 = (if n = 0 then {[]} else {})"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2123
  by (auto simp add: natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2124
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2125
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
  2126
  apply (auto simp add: set_replicate_conv_if natpermute_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2127
  apply (rule nth_equalityI)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2128
  apply simp_all
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2129
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2130
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2131
lemma natpermute_finite: "finite (natpermute n k)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2132
proof (induct k arbitrary: n)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2133
  case 0
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2134
  then show ?case
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2135
    apply (subst natpermute_split[of 0 0, simplified])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2136
    apply (simp add: natpermute_0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2137
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2138
next
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2139
  case (Suc k)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2140
  then show ?case unfolding natpermute_split [of k "Suc k", simplified]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2141
    apply -
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2142
    apply (rule finite_UN_I)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2143
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2144
    unfolding One_nat_def[symmetric] natlist_trivial_1
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2145
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2146
    done
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_contain_maximal:
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2150
  "{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
  2151
  (is "?A = ?B")
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2152
proof
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2153
  show "?A \<subseteq> ?B"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2154
  proof
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2155
    fix xs
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2156
    assume "xs \<in> ?A"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2157
    then have H: "xs \<in> natpermute n (k + 1)" and n: "n \<in> set xs"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2158
      by blast+
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2159
    then obtain i where i: "i \<in> {0.. k}" "xs!i = n"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2160
      unfolding in_set_conv_nth by (auto simp add: less_Suc_eq_le natpermute_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2161
    have eqs: "({0..k} - {i}) \<union> {i} = {0..k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2162
      using i by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2163
    have f: "finite({0..k} - {i})" "finite {i}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2164
      by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2165
    have d: "({0..k} - {i}) \<inter> {i} = {}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2166
      using i by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2167
    from H have "n = setsum (nth xs) {0..k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2168
      apply (simp add: natpermute_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2169
      apply (auto simp add: atLeastLessThanSuc_atLeastAtMost listsum_setsum_nth)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2170
      done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2171
    also have "\<dots> = n + setsum (nth xs) ({0..k} - {i})"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2172
      unfolding setsum.union_disjoint[OF f d, unfolded eqs] using i by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2173
    finally have zxs: "\<forall> j\<in> {0..k} - {i}. xs!j = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2174
      by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2175
    from H have xsl: "length xs = k+1"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2176
      by (simp add: natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2177
    from i have i': "i < length (replicate (k+1) 0)"   "i < k+1"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2178
      unfolding length_replicate by presburger+
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2179
    have "xs = replicate (k+1) 0 [i := n]"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2180
      apply (rule nth_equalityI)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2181
      unfolding xsl length_list_update length_replicate
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2182
      apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2183
      apply clarify
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2184
      unfolding nth_list_update[OF i'(1)]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2185
      using i zxs
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2186
      apply (case_tac "ia = i")
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2187
      apply (auto simp del: replicate.simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2188
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2189
    then show "xs \<in> ?B" using i by blast
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2190
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2191
  show "?B \<subseteq> ?A"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2192
  proof
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2193
    fix xs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2194
    assume "xs \<in> ?B"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2195
    then obtain i where i: "i \<in> {0..k}" and xs: "xs = replicate (k + 1) 0 [i:=n]"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2196
      by auto
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2197
    have nxs: "n \<in> set xs"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2198
      unfolding xs
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2199
      apply (rule set_update_memI)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2200
      using i apply simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2201
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2202
    have xsl: "length xs = k + 1"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2203
      by (simp only: xs length_replicate length_list_update)
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2204
    have "listsum xs = setsum (nth xs) {0..<k+1}"
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2205
      unfolding listsum_setsum_nth xsl ..
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2206
    also have "\<dots> = setsum (\<lambda>j. if j = i then n else 0) {0..< k+1}"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2207
      by (rule setsum.cong) (simp_all add: xs del: replicate.simps)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2208
    also have "\<dots> = n" using i by (simp add: setsum.delta)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2209
    finally have "xs \<in> natpermute n (k + 1)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2210
      using xsl unfolding natpermute_def mem_Collect_eq by blast
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2211
    then show "xs \<in> ?A"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2212
      using nxs by blast
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2213
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2214
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2215
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2216
text \<open>The general form.\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2217
lemma fps_setprod_nth:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2218
  fixes m :: nat
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2219
    and a :: "nat \<Rightarrow> 'a::comm_ring_1 fps"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2220
  shows "(setprod a {0 .. m}) $ n =
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2221
    setsum (\<lambda>v. setprod (\<lambda>j. (a j) $ (v!j)) {0..m}) (natpermute n (m+1))"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2222
  (is "?P m n")
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2223
proof (induct m arbitrary: n rule: nat_less_induct)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2224
  fix m n assume H: "\<forall>m' < m. \<forall>n. ?P m' n"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2225
  show "?P m n"
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2226
  proof (cases m)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2227
    case 0
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2228
    then show ?thesis
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2229
      apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2230
      unfolding natlist_trivial_1[where n = n, unfolded One_nat_def]
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2231
      apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2232
      done
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2233
  next
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2234
    case (Suc k)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2235
    then have km: "k < m" by arith
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2236
    have u0: "{0 .. k} \<union> {m} = {0..m}"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2237
      using Suc by (simp add: set_eq_iff) presburger
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2238
    have f0: "finite {0 .. k}" "finite {m}" by auto
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2239
    have d0: "{0 .. k} \<inter> {m} = {}" using Suc by auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2240
    have "(setprod a {0 .. m}) $ n = (setprod a {0 .. k} * a m) $ n"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2241
      unfolding setprod.union_disjoint[OF f0 d0, unfolded u0] by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2242
    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
  2243
      unfolding fps_mult_nth H[rule_format, OF km] ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2244
    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
  2245
      apply (simp add: Suc)
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2246
      unfolding natpermute_split[of m "m + 1", simplified, of n,
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2247
        unfolded natlist_trivial_1[unfolded One_nat_def] Suc]
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2248
      apply (subst setsum.UNION_disjoint)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2249
      apply simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2250
      apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2251
      unfolding image_Collect[symmetric]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2252
      apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2253
      apply (rule finite_imageI)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2254
      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
  2255
      apply (clarsimp simp add: set_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2256
      apply auto
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2257
      apply (rule setsum.cong)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2258
      apply (rule refl)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2259
      unfolding setsum_left_distrib
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2260
      apply (rule sym)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2261
      apply (rule_tac l = "\<lambda>xs. xs @ [n - x]" in setsum.reindex_cong)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2262
      apply (simp add: inj_on_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2263
      apply auto
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2264
      unfolding setprod.union_disjoint[OF f0 d0, unfolded u0, unfolded Suc]
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2265
      apply (clarsimp simp add: natpermute_def nth_append)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2266
      done
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2267
    finally show ?thesis .
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2268
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2269
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2270
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2271
text \<open>The special form for powers.\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2272
lemma fps_power_nth_Suc:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2273
  fixes m :: nat
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2274
    and a :: "'a::comm_ring_1 fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2275
  shows "(a ^ Suc m)$n = setsum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m}) (natpermute n (m+1))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2276
proof -
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2277
  have th0: "a^Suc m = setprod (\<lambda>i. a) {0..m}"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2278
    by (simp add: setprod_constant)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2279
  show ?thesis unfolding th0 fps_setprod_nth ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2280
qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2281
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2282
lemma fps_power_nth:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2283
  fixes m :: nat
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2284
    and a :: "'a::comm_ring_1 fps"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2285
  shows "(a ^m)$n =
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2286
    (if m=0 then 1$n else setsum (\<lambda>v. setprod (\<lambda>j. a $ (v!j)) {0..m - 1}) (natpermute n m))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2287
  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
  2288
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2289
lemma fps_nth_power_0:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2290
  fixes m :: nat
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2291
    and a :: "'a::comm_ring_1 fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2292
  shows "(a ^m)$0 = (a$0) ^ m"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2293
proof (cases m)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2294
  case 0
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2295
  then show ?thesis by simp
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2296
next
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2297
  case (Suc n)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2298
  then have c: "m = card {0..n}" by simp
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2299
  have "(a ^m)$0 = setprod (\<lambda>i. a$0) {0..n}"
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2300
    by (simp add: Suc fps_power_nth del: replicate.simps power_Suc)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2301
  also have "\<dots> = (a$0) ^ m"
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2302
   unfolding c by (rule setprod_constant) simp
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  2303
 finally show ?thesis .
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2304
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2305
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2306
lemma fps_compose_inj_right:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2307
  assumes a0: "a$0 = (0::'a::idom)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2308
    and a1: "a$1 \<noteq> 0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2309
  shows "(b oo a = c oo a) \<longleftrightarrow> b = c"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2310
  (is "?lhs \<longleftrightarrow>?rhs")
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2311
proof
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2312
  show ?lhs if ?rhs using that by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2313
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2314
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2315
    have "b$n = c$n" for n
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2316
    proof (induct n rule: nat_less_induct)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2317
      fix n
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2318
      assume H: "\<forall>m<n. b$m = c$m"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2319
      show "b$n = c$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2320
      proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2321
        case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2322
        from \<open>?lhs\<close> have "(b oo a)$n = (c oo a)$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2323
          by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2324
        then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2325
          using 0 by (simp add: fps_compose_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2326
      next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2327
        case (Suc n1)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2328
        have f: "finite {0 .. n1}" "finite {n}" by simp_all
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2329
        have eq: "{0 .. n1} \<union> {n} = {0 .. n}" using Suc by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2330
        have d: "{0 .. n1} \<inter> {n} = {}" using Suc by auto
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2331
        have seq: "(\<Sum>i = 0..n1. b $ i * a ^ i $ n) = (\<Sum>i = 0..n1. c $ i * a ^ i $ n)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2332
          apply (rule setsum.cong)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2333
          using H Suc
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2334
          apply auto
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2335
          done
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2336
        have th0: "(b oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + b$n * (a$1)^n"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2337
          unfolding fps_compose_nth setsum.union_disjoint[OF f d, unfolded eq] seq
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2338
          using startsby_zero_power_nth_same[OF a0]
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2339
          by simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2340
        have th1: "(c oo a) $n = (\<Sum>i = 0..n1. c $ i * a ^ i $ n) + c$n * (a$1)^n"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2341
          unfolding fps_compose_nth setsum.union_disjoint[OF f d, unfolded eq]
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2342
          using startsby_zero_power_nth_same[OF a0]
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2343
          by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2344
        from \<open>?lhs\<close>[unfolded fps_eq_iff, rule_format, of n] th0 th1 a1
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2345
        show ?thesis by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2346
      qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2347
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2348
    then show ?rhs by (simp add: fps_eq_iff)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2349
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2350
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2351
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2352
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2353
subsection \<open>Radicals\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2354
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2355
declare setprod.cong [fundef_cong]
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2356
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2357
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
  2358
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2359
  "radical r 0 a 0 = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2360
| "radical r 0 a (Suc n) = 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2361
| "radical r (Suc k) a 0 = r (Suc k) (a$0)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2362
| "radical r (Suc k) a (Suc n) =
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2363
    (a$ Suc n - setsum (\<lambda>xs. setprod (\<lambda>j. radical r (Suc k) a (xs ! j)) {0..k})
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2364
      {xs. xs \<in> natpermute (Suc n) (Suc k) \<and> Suc n \<notin> set xs}) /
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2365
    (of_nat (Suc k) * (radical r (Suc k) a 0)^k)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2366
  by pat_completeness auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2367
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2368
termination radical
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2369
proof
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2370
  let ?R = "measure (\<lambda>(r, k, a, n). n)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2371
  {
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2372
    show "wf ?R" by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2373
  next
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2374
    fix r k a n xs i
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2375
    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
  2376
    have False if c: "Suc n \<le> xs ! i"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2377
    proof -
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2378
      from xs i have "xs !i \<noteq> Suc n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2379
        by (auto simp add: in_set_conv_nth natpermute_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2380
      with c have c': "Suc n < xs!i" by arith
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2381
      have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2382
        by simp_all
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2383
      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
  2384
        by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2385
      have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2386
        using i by auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2387
      from xs have "Suc n = listsum xs"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2388
        by (simp add: natpermute_def)
46131
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2389
      also have "\<dots> = setsum (nth xs) {0..<Suc k}" using xs
ab07a3ef821c prefer listsum over foldl plus 0
haftmann
parents: 44174
diff changeset
  2390
        by (simp add: natpermute_def listsum_setsum_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2391
      also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2392
        unfolding eqs  setsum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2393
        unfolding setsum.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
  2394
        by simp
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2395
      finally show ?thesis using c' by simp
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2396
    qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2397
    then show "((r, Suc k, a, xs!i), r, Suc k, a, Suc n) \<in> ?R"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2398
      apply auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2399
      apply (metis not_less)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2400
      done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2401
  next
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2402
    fix r k a n
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2403
    show "((r, Suc k, a, 0), r, Suc k, a, Suc n) \<in> ?R" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2404
  }
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2405
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2406
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2407
definition "fps_radical r n a = Abs_fps (radical r n a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2408
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2409
lemma fps_radical0[simp]: "fps_radical r 0 a = 1"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2410
  apply (auto simp add: fps_eq_iff fps_radical_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2411
  apply (case_tac n)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2412
  apply auto
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2413
  done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2414
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2415
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
  2416
  by (cases n) (simp_all add: fps_radical_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2417
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2418
lemma fps_radical_power_nth[simp]:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2419
  assumes r: "(r k (a$0)) ^ k = a$0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2420
  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
  2421
proof (cases k)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2422
  case 0
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2423
  then show ?thesis by simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2424
next
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2425
  case (Suc h)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2426
  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
  2427
    unfolding fps_power_nth Suc by simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2428
  also have "\<dots> = (\<Prod>j\<in>{0..h}. r k (a$0))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2429
    apply (rule setprod.cong)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2430
    apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2431
    using Suc
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2432
    apply (subgoal_tac "replicate k 0 ! x = 0")
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2433
    apply (auto intro: nth_replicate simp del: replicate.simps)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2434
    done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2435
  also have "\<dots> = a$0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2436
    using r Suc by (simp add: setprod_constant)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2437
  finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2438
    using Suc by simp
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2439
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2440
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2441
lemma natpermute_max_card:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2442
  assumes n0: "n \<noteq> 0"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2443
  shows "card {xs \<in> natpermute n (k + 1). n \<in> set xs} = k + 1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2444
  unfolding natpermute_contain_maximal
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2445
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2446
  let ?A = "\<lambda>i. {replicate (k + 1) 0[i := n]}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2447
  let ?K = "{0 ..k}"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2448
  have fK: "finite ?K"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2449
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2450
  have fAK: "\<forall>i\<in>?K. finite (?A i)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2451
    by auto
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2452
  have d: "\<forall>i\<in> ?K. \<forall>j\<in> ?K. i \<noteq> j \<longrightarrow>
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2453
    {replicate (k + 1) 0[i := n]} \<inter> {replicate (k + 1) 0[j := n]} = {}"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2454
  proof clarify
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2455
    fix i j
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2456
    assume i: "i \<in> ?K" and j: "j \<in> ?K" and ij: "i \<noteq> j"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2457
    have False if eq: "replicate (k+1) 0 [i:=n] = replicate (k+1) 0 [j:= n]"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2458
    proof -
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2459
      have "(replicate (k+1) 0 [i:=n] ! i) = n"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2460
        using i by (simp del: replicate.simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2461
      moreover
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2462
      have "(replicate (k+1) 0 [j:=n] ! i) = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2463
        using i ij by (simp del: replicate.simps)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2464
      ultimately show ?thesis
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2465
        using eq n0 by (simp del: replicate.simps)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2466
    qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2467
    then show "{replicate (k + 1) 0[i := n]} \<inter> {replicate (k + 1) 0[j := n]} = {}"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2468
      by auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2469
  qed
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2470
  from card_UN_disjoint[OF fK fAK d]
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2471
  show "card (\<Union>i\<in>{0..k}. {replicate (k + 1) 0[i := n]}) = k + 1"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2472
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2473
qed
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2474
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2475
lemma power_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2476
  fixes a:: "'a::field_char_0 fps"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2477
  assumes a0: "a$0 \<noteq> 0"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2478
  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
  2479
    (is "?lhs \<longleftrightarrow> ?rhs")
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2480
proof
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2481
  let ?r = "fps_radical r (Suc k) a"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2482
  show ?rhs if r0: ?lhs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2483
  proof -
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2484
    from a0 r0 have r00: "r (Suc k) (a$0) \<noteq> 0" by auto
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2485
    have "?r ^ Suc k $ z = a$z" for z
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2486
    proof (induct z rule: nat_less_induct)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2487
      fix n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2488
      assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2489
      show "?r ^ Suc k $ n = a $n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2490
      proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2491
        case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2492
        then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2493
          using fps_radical_power_nth[of r "Suc k" a, OF r0] by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2494
      next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2495
        case (Suc n1)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2496
        then have "n \<noteq> 0" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2497
        let ?Pnk = "natpermute n (k + 1)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2498
        let ?Pnkn = "{xs \<in> ?Pnk. n \<in> set xs}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2499
        let ?Pnknn = "{xs \<in> ?Pnk. n \<notin> set xs}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2500
        have eq: "?Pnkn \<union> ?Pnknn = ?Pnk" by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2501
        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2502
        have f: "finite ?Pnkn" "finite ?Pnknn"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2503
          using finite_Un[of ?Pnkn ?Pnknn, unfolded eq]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2504
          by (metis natpermute_finite)+
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2505
        let ?f = "\<lambda>v. \<Prod>j\<in>{0..k}. ?r $ v ! j"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2506
        have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2507
        proof (rule setsum.cong)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2508
          fix v assume v: "v \<in> {xs \<in> natpermute n (k + 1). n \<in> set xs}"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2509
          let ?ths = "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2510
            fps_radical r (Suc k) a $ n * r (Suc k) (a $ 0) ^ k"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2511
          from v obtain i where i: "i \<in> {0..k}" "v = replicate (k+1) 0 [i:= n]"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2512
            unfolding natpermute_contain_maximal by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2513
          have "(\<Prod>j\<in>{0..k}. fps_radical r (Suc k) a $ v ! j) =
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2514
              (\<Prod>j\<in>{0..k}. if j = i then fps_radical r (Suc k) a $ n else r (Suc k) (a$0))"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2515
            apply (rule setprod.cong, simp)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2516
            using i r0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2517
            apply (simp del: replicate.simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2518
            done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2519
          also have "\<dots> = (fps_radical r (Suc k) a $ n) * r (Suc k) (a$0) ^ k"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2520
            using i r0 by (simp add: setprod_gen_delta)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2521
          finally show ?ths .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2522
        qed rule
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2523
        then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2524
          by (simp add: natpermute_max_card[OF \<open>n \<noteq> 0\<close>, simplified])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2525
        also have "\<dots> = a$n - setsum ?f ?Pnknn"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2526
          unfolding Suc using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2527
        finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2528
        have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2529
          unfolding fps_power_nth_Suc setsum.union_disjoint[OF f d, unfolded eq] ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2530
        also have "\<dots> = a$n" unfolding fn by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2531
        finally show ?thesis .
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2532
      qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2533
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2534
    then show ?thesis using r0 by (simp add: fps_eq_iff)
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2535
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2536
  show ?lhs if ?rhs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2537
  proof -
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2538
    from that have "((fps_radical r (Suc k) a) ^ (Suc k))$0 = a$0"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2539
      by simp
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2540
    then show ?thesis
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2541
      unfolding fps_power_nth_Suc
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2542
      by (simp add: setprod_constant del: replicate.simps)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2543
  qed
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2544
qed
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2545
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2546
(*
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2547
lemma power_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2548
  fixes a:: "'a::field_char_0 fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2549
  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
  2550
  shows "(fps_radical r (Suc k) a) ^ (Suc k) = a"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2551
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2552
  let ?r = "fps_radical r (Suc k) a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2553
  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
  2554
  {fix z have "?r ^ Suc k $ z = a$z"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2555
    proof(induct z rule: nat_less_induct)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2556
      fix n assume H: "\<forall>m<n. ?r ^ Suc k $ m = a$m"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2557
      {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
  2558
          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
  2559
      moreover
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2560
      {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
  2561
        have fK: "finite {0..k}" by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2562
        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
  2563
        let ?Pnk = "natpermute n (k + 1)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2564
        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
  2565
        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
  2566
        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
  2567
        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2568
        have f: "finite ?Pnkn" "finite ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2569
          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
  2570
          by (metis natpermute_finite)+
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2571
        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
  2572
        have "setsum ?f ?Pnkn = setsum (\<lambda>v. ?r $ n * r (Suc k) (a $ 0) ^ k) ?Pnkn"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2573
        proof(rule setsum.cong2)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2574
          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
  2575
          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
  2576
          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
  2577
            unfolding natpermute_contain_maximal by auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2578
          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))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2579
            apply (rule setprod.cong, simp)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2580
            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
  2581
          also have "\<dots> = (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
  2582
            unfolding setprod_gen_delta[OF fK] using i r0 by simp
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2583
          finally show ?ths .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2584
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2585
        then have "setsum ?f ?Pnkn = of_nat (k+1) * ?r $ n * r (Suc k) (a $ 0) ^ k"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2586
          by (simp add: natpermute_max_card[OF nz, simplified])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2587
        also have "\<dots> = a$n - setsum ?f ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2588
          unfolding n1 using r00 a0 by (simp add: field_simps fps_radical_def del: of_nat_Suc )
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2589
        finally have fn: "setsum ?f ?Pnkn = a$n - setsum ?f ?Pnknn" .
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2590
        have "(?r ^ Suc k)$n = setsum ?f ?Pnkn + setsum ?f ?Pnknn"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2591
          unfolding fps_power_nth_Suc setsum.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
  2592
        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
  2593
        finally have "?r ^ Suc k $ n = a $n" .}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2594
      ultimately  show "?r ^ Suc k $ n = a $n" by (cases n, auto)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2595
  qed }
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2596
  then show ?thesis by (simp add: fps_eq_iff)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2597
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2598
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2599
*)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2600
lemma eq_divide_imp':
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2601
  fixes c :: "'a::field"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2602
  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
  2603
  by (simp add: field_simps)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2604
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2605
lemma radical_unique:
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2606
  assumes r0: "(r (Suc k) (b$0)) ^ Suc k = b$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2607
    and a0: "r (Suc k) (b$0 ::'a::field_char_0) = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2608
    and b0: "b$0 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2609
  shows "a^(Suc k) = b \<longleftrightarrow> a = fps_radical r (Suc k) b"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2610
    (is "?lhs \<longleftrightarrow> ?rhs" is "_ \<longleftrightarrow> a = ?r")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2611
proof
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2612
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2613
    using that using power_radical[OF b0, of r k, unfolded r0] by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2614
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2615
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2616
    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
  2617
    have ceq: "card {0..k} = Suc k" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2618
    from a0 have a0r0: "a$0 = ?r$0" by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2619
    have "a $ n = ?r $ n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2620
    proof (induct n rule: nat_less_induct)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2621
      fix n
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2622
      assume h: "\<forall>m<n. a$m = ?r $m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2623
      show "a$n = ?r $ n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2624
      proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2625
        case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2626
        then show ?thesis using a0 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2627
      next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2628
        case (Suc n1)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2629
        have fK: "finite {0..k}" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2630
        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
  2631
        let ?Pnk = "natpermute n (Suc k)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2632
        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
  2633
        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
  2634
        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
  2635
        have d: "?Pnkn \<inter> ?Pnknn = {}" by blast
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2636
        have f: "finite ?Pnkn" "finite ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2637
          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
  2638
          by (metis natpermute_finite)+
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2639
        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
  2640
        let ?g = "\<lambda>v. \<Prod>j\<in>{0..k}. a $ v ! j"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2641
        have "setsum ?g ?Pnkn = setsum (\<lambda>v. a $ n * (?r$0)^k) ?Pnkn"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2642
        proof (rule setsum.cong)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2643
          fix v
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2644
          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
  2645
          let ?ths = "(\<Prod>j\<in>{0..k}. a $ v ! j) = a $ n * (?r$0)^k"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2646
          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
  2647
            unfolding Suc_eq_plus1 natpermute_contain_maximal
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2648
            by (auto simp del: replicate.simps)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2649
          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))"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2650
            apply (rule setprod.cong, simp)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2651
            using i a0
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2652
            apply (simp del: replicate.simps)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2653
            done
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2654
          also have "\<dots> = a $ n * (?r $ 0)^k"
46757
ad878aff9c15 removing finiteness goals
bulwahn
parents: 46131
diff changeset
  2655
            using i by (simp add: setprod_gen_delta)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2656
          finally show ?ths .
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2657
        qed rule
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2658
        then have th0: "setsum ?g ?Pnkn = of_nat (k+1) * a $ n * (?r $ 0)^k"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2659
          by (simp add: natpermute_max_card[OF nz, simplified])
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2660
        have th1: "setsum ?g ?Pnknn = setsum ?f ?Pnknn"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2661
        proof (rule setsum.cong, rule refl, rule setprod.cong, simp)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2662
          fix xs i
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2663
          assume xs: "xs \<in> ?Pnknn" and i: "i \<in> {0..k}"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2664
          have False if c: "n \<le> xs ! i"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2665
          proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2666
            from xs i have "xs ! i \<noteq> n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2667
              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
  2668
            with c have c': "n < xs!i" by arith
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2669
            have fths: "finite {0 ..< i}" "finite {i}" "finite {i+1..<Suc k}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2670
              by simp_all
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2671
            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
  2672
              by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2673
            have eqs: "{0..<Suc k} = {0 ..< i} \<union> ({i} \<union> {i+1 ..< Suc k})"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2674
              using i by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2675
            from xs have "n = listsum xs"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2676
              by (simp add: natpermute_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2677
            also have "\<dots> = setsum (nth xs) {0..<Suc k}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2678
              using xs by (simp add: natpermute_def listsum_setsum_nth)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2679
            also have "\<dots> = xs!i + setsum (nth xs) {0..<i} + setsum (nth xs) {i+1..<Suc k}"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2680
              unfolding eqs  setsum.union_disjoint[OF fths(1) finite_UnI[OF fths(2,3)] d(1)]
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2681
              unfolding setsum.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
  2682
              by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2683
            finally show ?thesis using c' by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2684
          qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2685
          then have thn: "xs!i < n" by presburger
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2686
          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
  2687
        qed
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2688
        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
  2689
          by (simp add: field_simps del: of_nat_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2690
        from \<open>?lhs\<close> have "b$n = a^Suc k $ n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2691
          by (simp add: fps_eq_iff)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2692
        also have "a ^ Suc k$n = setsum ?g ?Pnkn + setsum ?g ?Pnknn"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2693
          unfolding fps_power_nth_Suc
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2694
          using setsum.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
  2695
            unfolded eq, of ?g] by simp
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2696
        also have "\<dots> = of_nat (k+1) * a $ n * (?r $ 0)^k + setsum ?f ?Pnknn"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2697
          unfolding th0 th1 ..
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2698
        finally have "of_nat (k+1) * a $ n * (?r $ 0)^k = b$n - setsum ?f ?Pnknn"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2699
          by simp
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2700
        then have "a$n = (b$n - setsum ?f ?Pnknn) / (of_nat (k+1) * (?r $ 0)^k)"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2701
          apply -
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2702
          apply (rule eq_divide_imp')
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2703
          using r00
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2704
          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
  2705
          apply (simp add: ac_simps)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2706
          done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2707
        then show ?thesis
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  2708
          apply (simp del: of_nat_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2709
          unfolding fps_radical_def Suc
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2710
          apply (simp add: field_simps Suc th00 del: of_nat_Suc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2711
          done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2712
      qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2713
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2714
    then show ?rhs by (simp add: fps_eq_iff)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2715
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2716
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2717
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2718
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2719
lemma radical_power:
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2720
  assumes r0: "r (Suc k) ((a$0) ^ Suc k) = a$0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2721
    and a0: "(a$0 :: 'a::field_char_0) \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2722
  shows "(fps_radical r (Suc k) (a ^ Suc k)) = a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2723
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2724
  let ?ak = "a^ Suc k"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2725
  have ak0: "?ak $ 0 = (a$0) ^ Suc k"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2726
    by (simp add: fps_nth_power_0 del: power_Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2727
  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
  2728
    using ak0 by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2729
  from r0 ak0 have th1: "r (Suc k) (a ^ Suc k $ 0) = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2730
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2731
  from ak0 a0 have ak00: "?ak $ 0 \<noteq>0 "
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2732
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2733
  from radical_unique[of r k ?ak a, OF th0 th1 ak00] show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2734
    by metis
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2735
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2736
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2737
lemma fps_deriv_radical:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2738
  fixes a :: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2739
  assumes r0: "(r (Suc k) (a$0)) ^ Suc k = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2740
    and a0: "a$0 \<noteq> 0"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2741
  shows "fps_deriv (fps_radical r (Suc k) a) =
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2742
    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
  2743
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2744
  let ?r = "fps_radical r (Suc k) a"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2745
  let ?w = "(fps_const (of_nat (Suc k)) * ?r ^ k)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2746
  from a0 r0 have r0': "r (Suc k) (a$0) \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2747
    by auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2748
  from r0' have w0: "?w $ 0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2749
    by (simp del: of_nat_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2750
  note th0 = inverse_mult_eq_1[OF w0]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2751
  let ?iw = "inverse ?w"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2752
  from iffD1[OF power_radical[of a r], OF a0 r0]
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2753
  have "fps_deriv (?r ^ Suc k) = fps_deriv a"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2754
    by simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2755
  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
  2756
    by (simp add: fps_deriv_power ac_simps del: power_Suc)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2757
  then have "?iw * fps_deriv ?r * ?w = ?iw * fps_deriv a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2758
    by simp
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  2759
  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
  2760
    by (subst fps_divide_unit) (auto simp del: of_nat_Suc)
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2761
  then show ?thesis unfolding th0 by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2762
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2763
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2764
lemma radical_mult_distrib:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2765
  fixes a :: "'a::field_char_0 fps"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2766
  assumes k: "k > 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2767
    and ra0: "r k (a $ 0) ^ k = a $ 0"
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2768
    and rb0: "r k (b $ 0) ^ k = b $ 0"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2769
    and a0: "a $ 0 \<noteq> 0"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2770
    and b0: "b $ 0 \<noteq> 0"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  2771
  shows "r k ((a * b) $ 0) = r k (a $ 0) * r k (b $ 0) \<longleftrightarrow>
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2772
    fps_radical r k (a * b) = fps_radical r k a * fps_radical r k b"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2773
    (is "?lhs \<longleftrightarrow> ?rhs")
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2774
proof
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2775
  show ?rhs if r0': ?lhs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2776
  proof -
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2777
    from r0' have r0: "(r k ((a * b) $ 0)) ^ k = (a * b) $ 0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2778
      by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2779
    show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2780
    proof (cases k)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2781
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2782
      then show ?thesis using r0' by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2783
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2784
      case (Suc h)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2785
      let ?ra = "fps_radical r (Suc h) a"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2786
      let ?rb = "fps_radical r (Suc h) b"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2787
      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
  2788
        using r0' Suc by (simp add: fps_mult_nth)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2789
      have ab0: "(a*b) $ 0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2790
        using a0 b0 by (simp add: fps_mult_nth)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2791
      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
  2792
        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
  2793
      show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2794
        by (auto simp add: power_mult_distrib simp del: power_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2795
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2796
  qed
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2797
  show ?lhs if ?rhs
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2798
  proof -
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2799
    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
  2800
      by simp
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2801
    then show ?thesis
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2802
      using k by (simp add: fps_mult_nth)
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2803
  qed
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2804
qed
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2805
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2806
(*
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2807
lemma radical_mult_distrib:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2808
  fixes a:: "'a::field_char_0 fps"
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2809
  assumes
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2810
  ra0: "r k (a $ 0) ^ k = a $ 0"
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2811
  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
  2812
  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
  2813
  and a0: "a$0 \<noteq> 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2814
  and b0: "b$0 \<noteq> 0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2815
  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
  2816
proof-
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2817
  from r0' have r0: "(r (k) ((a*b)$0)) ^ k = (a*b)$0"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2818
    by (simp add: fps_mult_nth ra0 rb0 power_mult_distrib)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  2819
  {assume "k=0" then have ?thesis by simp}
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2820
  moreover
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2821
  {fix h assume k: "k = Suc h"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2822
  let ?ra = "fps_radical r (Suc h) a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2823
  let ?rb = "fps_radical r (Suc h) b"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2824
  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
  2825
    using r0' k by (simp add: fps_mult_nth)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2826
  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
  2827
  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
  2828
    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
  2829
  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
  2830
ultimately show ?thesis by (cases k, auto)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2831
qed
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2832
*)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2833
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  2834
lemma fps_divide_1 [simp]: "(a :: 'a::field fps) / 1 = a"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  2835
  by (fact divide_1)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2836
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2837
lemma radical_divide:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2838
  fixes a :: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2839
  assumes kp: "k > 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2840
    and ra0: "(r k (a $ 0)) ^ k = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2841
    and rb0: "(r k (b $ 0)) ^ k = b $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2842
    and a0: "a$0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2843
    and b0: "b$0 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2844
  shows "r k ((a $ 0) / (b$0)) = r k (a$0) / r k (b $ 0) \<longleftrightarrow>
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2845
    fps_radical r k (a/b) = fps_radical r k a / fps_radical r k b"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2846
  (is "?lhs = ?rhs")
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2847
proof
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2848
  let ?r = "fps_radical r k"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2849
  from kp obtain h where k: "k = Suc h"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  2850
    by (cases k) auto
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2851
  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
  2852
  have rb0': "r k (b$0) \<noteq> 0" using b0 rb0 k by auto
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2853
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2854
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2855
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2856
    from that have "?r (a/b) $ 0 = (?r a / ?r b)$0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2857
      by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2858
    then show ?thesis
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  2859
      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
  2860
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2861
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2862
  proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  2863
    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
  2864
      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
  2865
    have th0: "r k ((a/b)$0) ^ k = (a/b)$0"
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
  2866
      by (simp add: \<open>?lhs\<close> power_divide ra0 rb0)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2867
    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
  2868
    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
  2869
      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
  2870
    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
  2871
      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
  2872
    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
  2873
    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
  2874
    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
  2875
      by (simp add: fps_divide_unit power_mult_distrib fps_inverse_power[symmetric])
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  2876
      
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  2877
    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
  2878
    show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2879
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2880
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2881
31073
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2882
lemma radical_inverse:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  2883
  fixes a :: "'a::field_char_0 fps"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2884
  assumes k: "k > 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2885
    and ra0: "r k (a $ 0) ^ k = a $ 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2886
    and r1: "(r k 1)^k = 1"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2887
    and a0: "a$0 \<noteq> 0"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2888
  shows "r k (inverse (a $ 0)) = r k 1 / (r k (a $ 0)) \<longleftrightarrow>
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  2889
    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
  2890
  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
  2891
  by (simp add: divide_inverse fps_divide_def)
4b44c4d08aa6 Generalized distributivity theorems of radicals over multiplication, division and inverses
chaieb
parents: 31021
diff changeset
  2892
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2893
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2894
subsection \<open>Derivative of composition\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2895
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  2896
lemma fps_compose_deriv:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2897
  fixes a :: "'a::idom fps"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2898
  assumes b0: "b$0 = 0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2899
  shows "fps_deriv (a oo b) = ((fps_deriv a) oo b) * fps_deriv b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2900
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2901
  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
  2902
  proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2903
    have "(fps_deriv (a oo b))$n = setsum (\<lambda>i. a $ i * (fps_deriv (b^i))$n) {0.. Suc n}"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  2904
      by (simp add: fps_compose_def field_simps setsum_right_distrib del: of_nat_Suc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2905
    also have "\<dots> = setsum (\<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
  2906
      by (simp add: field_simps fps_deriv_power del: fps_mult_left_const_nth of_nat_Suc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2907
    also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (((b^(i - 1)) * fps_deriv b))$n) {0.. Suc n}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2908
      unfolding fps_mult_left_const_nth  by (simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2909
    also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {0.. Suc n}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2910
      unfolding fps_mult_nth ..
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2911
    also have "\<dots> = setsum (\<lambda>i. of_nat i * a$i * (setsum (\<lambda>j. (b^ (i - 1))$j * (fps_deriv b)$(n - j)) {0..n})) {1.. Suc n}"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2912
      apply (rule setsum.mono_neutral_right)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2913
      apply (auto simp add: mult_delta_left setsum.delta not_le)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2914
      done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2915
    also have "\<dots> = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2916
      unfolding fps_deriv_nth
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  2917
      by (rule setsum.reindex_cong [of Suc]) (auto simp add: mult.assoc)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2918
    finally have th0: "(fps_deriv (a oo b))$n =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2919
      setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}" .
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2920
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2921
    have "(((fps_deriv a) oo b) * (fps_deriv b))$n = setsum (\<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
  2922
      unfolding fps_mult_nth by (simp add: ac_simps)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2923
    also have "\<dots> = setsum (\<lambda>i. setsum (\<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}"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  2924
      unfolding fps_deriv_nth fps_compose_nth setsum_right_distrib mult.assoc
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2925
      apply (rule setsum.cong)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2926
      apply (rule refl)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2927
      apply (rule setsum.mono_neutral_left)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2928
      apply (simp_all add: subset_eq)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2929
      apply clarify
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2930
      apply (subgoal_tac "b^i$x = 0")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2931
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2932
      apply (rule startsby_zero_power_prefix[OF b0, rule_format])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2933
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2934
      done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2935
    also have "\<dots> = setsum (\<lambda>i. of_nat (i + 1) * a$(i+1) * (setsum (\<lambda>j. (b^ i)$j * of_nat (n - j + 1) * b$(n - j + 1)) {0..n})) {0.. n}"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2936
      unfolding setsum_right_distrib
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2937
      apply (subst setsum.commute)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2938
      apply (rule setsum.cong, rule refl)+
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2939
      apply simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2940
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2941
    finally show ?thesis
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2942
      unfolding th0 by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2943
  qed
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2944
  then show ?thesis by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2945
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2946
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2947
lemma fps_mult_X_plus_1_nth:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2948
  "((1+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
  2949
proof (cases n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2950
  case 0
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2951
  then show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2952
    by (simp add: fps_mult_nth)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2953
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2954
  case (Suc m)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2955
  have "((1 + X)*a) $ n = setsum (\<lambda>i. (1 + X) $ i * a $ (n - i)) {0..n}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2956
    by (simp add: fps_mult_nth)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2957
  also have "\<dots> = setsum (\<lambda>i. (1+X)$i * a$(n-i)) {0.. 1}"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  2958
    unfolding Suc by (rule setsum.mono_neutral_right) auto
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2959
  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
  2960
    by (simp add: Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2961
  finally show ?thesis .
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2962
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2963
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2964
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2965
subsection \<open>Finite FPS (i.e. polynomials) and X\<close>
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2966
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2967
lemma fps_poly_sum_X:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2968
  assumes "\<forall>i > n. a$i = (0::'a::comm_ring_1)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2969
  shows "a = setsum (\<lambda>i. fps_const (a$i) * X^i) {0..n}" (is "a = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2970
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2971
  have "a$i = ?r$i" for i
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2972
    unfolding fps_setsum_nth fps_mult_left_const_nth X_power_nth
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2973
    by (simp add: mult_delta_right setsum.delta' assms)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2974
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2975
    unfolding fps_eq_iff by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2976
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2977
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2978
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2979
subsection \<open>Compositional inverses\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2980
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  2981
fun compinv :: "'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2982
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2983
  "compinv a 0 = X$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2984
| "compinv a (Suc n) =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2985
    (X$ Suc n - setsum (\<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
  2986
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2987
definition "fps_inv a = Abs_fps (compinv a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2988
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2989
lemma fps_inv:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2990
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2991
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2992
  shows "fps_inv a oo a = X"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2993
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  2994
  let ?i = "fps_inv a oo a"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2995
  have "?i $n = X$n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2996
  proof (induct n rule: nat_less_induct)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  2997
    fix n
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2998
    assume h: "\<forall>m<n. ?i$m = X$m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  2999
    show "?i $ n = X$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3000
    proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3001
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3002
      then show ?thesis using a0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3003
        by (simp add: fps_compose_nth fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3004
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3005
      case (Suc n1)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3006
      have "?i $ n = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} + fps_inv a $ Suc n1 * (a $ 1)^ Suc n1"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3007
        by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3008
      also have "\<dots> = setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1} +
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3009
        (X$ Suc n1 - setsum (\<lambda>i. (fps_inv a $ i) * (a^i)$n) {0 .. n1})"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3010
        using a0 a1 Suc by (simp add: fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3011
      also have "\<dots> = X$n" using Suc by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3012
      finally show ?thesis .
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3013
    qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3014
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3015
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3016
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3017
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3018
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3019
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3020
fun gcompinv :: "'a fps \<Rightarrow> 'a fps \<Rightarrow> nat \<Rightarrow> 'a::field"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3021
where
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3022
  "gcompinv b a 0 = b$0"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3023
| "gcompinv b a (Suc n) =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3024
    (b$ Suc n - setsum (\<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
  3025
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3026
definition "fps_ginv b a = Abs_fps (gcompinv b a)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3027
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3028
lemma fps_ginv:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3029
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3030
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3031
  shows "fps_ginv b a oo a = b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3032
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3033
  let ?i = "fps_ginv b a oo a"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3034
  have "?i $n = b$n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3035
  proof (induct n rule: nat_less_induct)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3036
    fix n
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3037
    assume h: "\<forall>m<n. ?i$m = b$m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3038
    show "?i $ n = b$n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3039
    proof (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3040
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3041
      then show ?thesis using a0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3042
        by (simp add: fps_compose_nth fps_ginv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3043
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3044
      case (Suc n1)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3045
      have "?i $ n = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} + fps_ginv b a $ Suc n1 * (a $ 1)^ Suc n1"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3046
        by (simp only: fps_compose_nth) (simp add: Suc startsby_zero_power_nth_same [OF a0] del: power_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3047
      also have "\<dots> = setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1} +
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3048
        (b$ Suc n1 - setsum (\<lambda>i. (fps_ginv b a $ i) * (a^i)$n) {0 .. n1})"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3049
        using a0 a1 Suc by (simp add: fps_ginv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3050
      also have "\<dots> = b$n" using Suc by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3051
      finally show ?thesis .
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3052
    qed
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3053
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3054
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3055
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3056
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3057
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3058
lemma fps_inv_ginv: "fps_inv = fps_ginv 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
  3059
  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
  3060
  apply (induct_tac n rule: nat_less_induct)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  3061
  apply auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3062
  apply (case_tac na)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3063
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3064
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3065
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3066
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3067
lemma fps_compose_1[simp]: "1 oo a = 1"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3068
  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3069
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3070
lemma fps_compose_0[simp]: "0 oo a = 0"
29913
89eadbe71e97 add mult_delta lemmas; simplify some proofs
huffman
parents: 29912
diff changeset
  3071
  by (simp add: fps_eq_iff fps_compose_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3072
60867
86e7560e07d0 slight cleanup of lemmas
haftmann
parents: 60679
diff changeset
  3073
lemma fps_compose_0_right[simp]: "a oo 0 = fps_const (a $ 0)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3074
  by (auto simp add: fps_eq_iff fps_compose_nth power_0_left setsum.neutral)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3075
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3076
lemma fps_compose_add_distrib: "(a + b) oo c = (a oo c) + (b oo c)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3077
  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum.distrib)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3078
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3079
lemma fps_compose_setsum_distrib: "(setsum f S) oo a = setsum (\<lambda>i. f i oo a) S"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3080
proof (cases "finite S")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3081
  case True
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3082
  show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3083
  proof (rule finite_induct[OF True])
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3084
    show "setsum f {} oo a = (\<Sum>i\<in>{}. f i oo a)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3085
      by simp
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3086
  next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3087
    fix x F
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3088
    assume fF: "finite F"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3089
      and xF: "x \<notin> F"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3090
      and h: "setsum f F oo a = setsum (\<lambda>i. f i oo a) F"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3091
    show "setsum f (insert x F) oo a  = setsum (\<lambda>i. f i oo a) (insert x F)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3092
      using fF xF h by (simp add: fps_compose_add_distrib)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3093
  qed
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3094
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3095
  case False
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3096
  then show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3097
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3098
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3099
lemma convolution_eq:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3100
  "setsum (\<lambda>i. a (i :: nat) * b (n - i)) {0 .. n} =
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3101
    setsum (\<lambda>(i,j). a i * b j) {(i,j). i \<le> n \<and> j \<le> n \<and> i + j = n}"
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  3102
  by (rule setsum.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
  3103
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3104
lemma product_composition_lemma:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3105
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3106
    and d0: "d$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3107
  shows "((a oo c) * (b oo d))$n =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3108
    setsum (\<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
  3109
proof -
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3110
  let ?S = "{(k::nat, m::nat). k + m \<le> n}"
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3111
  have s: "?S \<subseteq> {0..n} <*> {0..n}" by (auto simp add: subset_eq)
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3112
  have f: "finite {(k::nat, m::nat). k + m \<le> n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3113
    apply (rule finite_subset[OF s])
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3114
    apply auto
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3115
    done
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3116
  have "?r =  setsum (\<lambda>i. setsum (\<lambda>(k,m). a$k * (c^k)$i * b$m * (d^m) $ (n - i)) {(k,m). k + m \<le> n}) {0..n}"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3117
    apply (simp add: fps_mult_nth setsum_right_distrib)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3118
    apply (subst setsum.commute)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3119
    apply (rule setsum.cong)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3120
    apply (auto simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3121
    done
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3122
  also have "\<dots> = ?l"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3123
    apply (simp add: fps_mult_nth fps_compose_nth setsum_product)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3124
    apply (rule setsum.cong)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3125
    apply (rule refl)
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  3126
    apply (simp add: setsum.cartesian_product mult.assoc)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3127
    apply (rule setsum.mono_neutral_right[OF f])
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3128
    apply (simp add: subset_eq)
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3129
    apply presburger
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3130
    apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3131
    apply (rule ccontr)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3132
    apply (clarsimp simp add: not_le)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3133
    apply (case_tac "x < aa")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3134
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3135
    apply (frule_tac startsby_zero_power_prefix[rule_format, OF c0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3136
    apply blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3137
    apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3138
    apply (frule_tac startsby_zero_power_prefix[rule_format, OF d0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3139
    apply blast
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3140
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3141
  finally show ?thesis by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3142
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3143
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3144
lemma product_composition_lemma':
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3145
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3146
    and d0: "d$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3147
  shows "((a oo c) * (b oo d))$n =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3148
    setsum (\<lambda>k. setsum (\<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
  3149
  unfolding product_composition_lemma[OF c0 d0]
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3150
  unfolding setsum.cartesian_product
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3151
  apply (rule setsum.mono_neutral_left)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3152
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3153
  apply (clarsimp simp add: subset_eq)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3154
  apply clarsimp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3155
  apply (rule ccontr)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3156
  apply (subgoal_tac "(c^aa * d^ba) $ n = 0")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3157
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3158
  unfolding fps_mult_nth
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3159
  apply (rule setsum.neutral)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3160
  apply (clarsimp simp add: not_le)
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51107
diff changeset
  3161
  apply (case_tac "x < aa")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3162
  apply (rule startsby_zero_power_prefix[OF c0, rule_format])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3163
  apply simp
51489
f738e6dbd844 fundamental revision of big operators on sets
haftmann
parents: 51107
diff changeset
  3164
  apply (subgoal_tac "n - x < ba")
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3165
  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
  3166
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3167
  apply arith
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3168
  done
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3169
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3170
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3171
lemma setsum_pair_less_iff:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3172
  "setsum (\<lambda>((k::nat),m). a k * b m * c (k + m)) {(k,m). k + m \<le> n} =
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3173
    setsum (\<lambda>s. setsum (\<lambda>i. a i * b (s - i) * c s) {0..s}) {0..n}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3174
  (is "?l = ?r")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3175
proof -
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3176
  let ?KM = "{(k,m). k + m \<le> n}"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3177
  let ?f = "\<lambda>s. UNION {(0::nat)..s} (\<lambda>i. {(i,s - i)})"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3178
  have th0: "?KM = UNION {0..n} ?f"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59815
diff changeset
  3179
    by (auto simp add: set_eq_iff Bex_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3180
  show "?l = ?r "
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3181
    unfolding th0
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3182
    apply (subst setsum.UNION_disjoint)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3183
    apply auto
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3184
    apply (subst setsum.UNION_disjoint)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3185
    apply auto
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3186
    done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3187
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3188
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3189
lemma fps_compose_mult_distrib_lemma:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3190
  assumes c0: "c$0 = (0::'a::idom)"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3191
  shows "((a oo c) * (b oo c))$n = setsum (\<lambda>s. setsum (\<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
  3192
  unfolding product_composition_lemma[OF c0 c0] power_add[symmetric]
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3193
  unfolding setsum_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
  3194
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3195
lemma fps_compose_mult_distrib:
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54452
diff changeset
  3196
  assumes c0: "c $ 0 = (0::'a::idom)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54452
diff changeset
  3197
  shows "(a * b) oo c = (a oo c) * (b oo c)"
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54452
diff changeset
  3198
  apply (simp add: fps_eq_iff fps_compose_mult_distrib_lemma [OF c0])
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3199
  apply (simp add: fps_compose_nth fps_mult_nth setsum_left_distrib)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3200
  done
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3201
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3202
lemma fps_compose_setprod_distrib:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3203
  assumes c0: "c$0 = (0::'a::idom)"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3204
  shows "setprod a S oo c = setprod (\<lambda>k. a k oo c) S"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3205
  apply (cases "finite S")
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3206
  apply simp_all
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3207
  apply (induct S rule: finite_induct)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3208
  apply simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3209
  apply (simp add: fps_compose_mult_distrib[OF c0])
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3210
  done
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3211
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3212
lemma fps_compose_power:
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3213
  assumes c0: "c$0 = (0::'a::idom)"
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3214
  shows "(a oo c)^n = a^n oo c"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3215
proof (cases n)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3216
  case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3217
  then show ?thesis by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3218
next
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3219
  case (Suc m)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3220
  have th0: "a^n = setprod (\<lambda>k. a) {0..m}" "(a oo c) ^ n = setprod (\<lambda>k. a oo c) {0..m}"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3221
    by (simp_all add: setprod_constant Suc)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3222
  then show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3223
    by (simp add: fps_compose_setprod_distrib[OF c0])
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3224
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3225
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3226
lemma fps_compose_uminus: "- (a::'a::ring_1 fps) oo c = - (a oo c)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3227
  by (simp add: fps_eq_iff fps_compose_nth field_simps setsum_negf[symmetric])
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3228
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3229
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
  3230
  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
  3231
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3232
lemma X_fps_compose: "X oo a = Abs_fps (\<lambda>n. if n = 0 then (0::'a::comm_ring_1) else a$n)"
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3233
  by (simp add: fps_eq_iff fps_compose_nth mult_delta_left setsum.delta)
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3234
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3235
lemma fps_inverse_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3236
  assumes b0: "(b$0 :: 'a::field) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3237
    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
  3238
  shows "inverse a oo b = inverse (a oo b)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3239
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3240
  let ?ia = "inverse a"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3241
  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
  3242
  let ?iab = "inverse ?ab"
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3243
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3244
  from a0 have ia0: "?ia $ 0 \<noteq> 0" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3245
  from a0 have ab0: "?ab $ 0 \<noteq> 0" by (simp add: fps_compose_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3246
  have "(?ia oo b) *  (a oo b) = 1"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3247
    unfolding fps_compose_mult_distrib[OF b0, symmetric]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3248
    unfolding inverse_mult_eq_1[OF a0]
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3249
    fps_compose_1 ..
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3250
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3251
  then have "(?ia oo b) *  (a oo b) * ?iab  = 1 * ?iab" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3252
  then have "(?ia oo b) *  (?iab * (a oo b))  = ?iab" by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3253
  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
  3254
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3255
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3256
lemma fps_divide_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3257
  assumes c0: "(c$0 :: 'a::field) = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3258
    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
  3259
  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
  3260
    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
  3261
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3262
lemma gp:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3263
  assumes a0: "a$0 = (0::'a::field)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3264
  shows "(Abs_fps (\<lambda>n. 1)) oo a = 1/(1 - a)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3265
    (is "?one oo a = _")
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3266
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3267
  have o0: "?one $ 0 \<noteq> 0" by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3268
  have th0: "(1 - 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
  3269
  from fps_inverse_gp[where ?'a = 'a]
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3270
  have "inverse ?one = 1 - X" by (simp add: fps_eq_iff)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3271
  then have "inverse (inverse ?one) = inverse (1 - X)" by simp
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3272
  then have th: "?one = 1/(1 - 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
  3273
    by (simp add: fps_divide_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3274
  show ?thesis
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3275
    unfolding th
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3276
    unfolding fps_divide_compose[OF a0 th0]
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3277
    fps_compose_1 fps_compose_sub_distrib X_fps_compose_startby0[OF a0] ..
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3278
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3279
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3280
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
  3281
  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
  3282
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3283
lemma fps_compose_radical:
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3284
  assumes b0: "b$0 = (0::'a::field_char_0)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3285
    and ra0: "r (Suc k) (a$0) ^ Suc k = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3286
    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
  3287
  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
  3288
proof -
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3289
  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
  3290
  let ?ab = "a oo b"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3291
  have ab0: "?ab $ 0 = a$0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3292
    by (simp add: fps_compose_def)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3293
  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
  3294
    by simp_all
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3295
  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
  3296
    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
  3297
  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
  3298
    unfolding fps_compose_power[OF b0]
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3299
    unfolding iffD1[OF power_radical[of a r k], OF a0 ra0]  ..
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3300
  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
  3301
  show ?thesis  .
31199
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3302
qed
10d413b08fa7 FPS composition distributes over inverses, division and arbitrary nth roots. General geometric series theorem
chaieb
parents: 31148
diff changeset
  3303
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3304
lemma fps_const_mult_apply_left: "fps_const c * (a oo b) = (fps_const c * a) oo b"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  3305
  by (simp add: fps_eq_iff fps_compose_nth setsum_right_distrib mult.assoc)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3306
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3307
lemma fps_const_mult_apply_right:
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3308
  "(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
  3309
  by (auto simp add: fps_const_mult_apply_left mult.commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3310
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3311
lemma fps_compose_assoc:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3312
  assumes c0: "c$0 = (0::'a::idom)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3313
    and b0: "b$0 = 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3314
  shows "a oo (b oo c) = a oo b oo c" (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3315
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3316
  have "?l$n = ?r$n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3317
  proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3318
    have "?l$n = (setsum (\<lambda>i. (fps_const (a$i) * b^i) oo c) {0..n})$n"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3319
      by (simp add: fps_compose_nth fps_compose_power[OF c0] fps_const_mult_apply_left
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  3320
        setsum_right_distrib mult.assoc fps_setsum_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3321
    also have "\<dots> = ((setsum (\<lambda>i. fps_const (a$i) * b^i) {0..n}) oo c)$n"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3322
      by (simp add: fps_compose_setsum_distrib)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3323
    also have "\<dots> = ?r$n"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  3324
      apply (simp add: fps_compose_nth fps_setsum_nth setsum_left_distrib mult.assoc)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3325
      apply (rule setsum.cong)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3326
      apply (rule refl)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3327
      apply (rule setsum.mono_neutral_right)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3328
      apply (auto simp add: not_le)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3329
      apply (erule startsby_zero_power_prefix[OF b0, rule_format])
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3330
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3331
    finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3332
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3333
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3334
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3335
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3336
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3337
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3338
lemma fps_X_power_compose:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3339
  assumes a0: "a$0=0"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3340
  shows "X^k oo a = (a::'a::idom fps)^k"
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3341
  (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3342
proof (cases k)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3343
  case 0
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3344
  then show ?thesis by simp
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3345
next
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3346
  case (Suc h)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3347
  have "?l $ n = ?r $n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3348
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3349
    consider "k > n" | "k \<le> n" by arith
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3350
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3351
    proof cases
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3352
      case 1
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3353
      then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3354
        using a0 startsby_zero_power_prefix[OF a0] Suc
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3355
        by (simp add: fps_compose_nth del: power_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3356
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3357
      case 2
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3358
      then show ?thesis
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3359
        by (simp add: fps_compose_nth mult_delta_left setsum.delta)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3360
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3361
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3362
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3363
    unfolding fps_eq_iff by blast
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3364
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3365
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3366
lemma fps_inv_right:
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3367
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3368
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3369
  shows "a oo fps_inv a = X"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3370
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3371
  let ?ia = "fps_inv a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3372
  let ?iaa = "a oo fps_inv a"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3373
  have th0: "?ia $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3374
    by (simp add: fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3375
  have th1: "?iaa $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3376
    using a0 a1 by (simp add: fps_inv_def fps_compose_nth)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3377
  have th2: "X$0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3378
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3379
  from fps_inv[OF a0 a1] have "a oo (fps_inv a oo a) = a oo X"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3380
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3381
  then have "(a oo fps_inv a) oo a = X oo a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3382
    by (simp add: fps_compose_assoc[OF a0 th0] X_fps_compose_startby0[OF a0])
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3383
  with fps_compose_inj_right[OF a0 a1] show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3384
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3385
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3386
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3387
lemma fps_inv_deriv:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3388
  assumes a0: "a$0 = (0::'a::field)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3389
    and a1: "a$1 \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3390
  shows "fps_deriv (fps_inv a) = inverse (fps_deriv a oo fps_inv a)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3391
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3392
  let ?ia = "fps_inv a"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3393
  let ?d = "fps_deriv a oo ?ia"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3394
  let ?dia = "fps_deriv ?ia"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3395
  have ia0: "?ia$0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3396
    by (simp add: fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3397
  have th0: "?d$0 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3398
    using a1 by (simp add: fps_compose_nth)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3399
  from fps_inv_right[OF a0 a1] have "?d * ?dia = 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3400
    by (simp add: fps_compose_deriv[OF ia0, of a, symmetric] )
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3401
  then have "inverse ?d * ?d * ?dia = inverse ?d * 1"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3402
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3403
  with inverse_mult_eq_1 [OF th0] show "?dia = inverse ?d"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3404
    by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3405
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3406
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3407
lemma fps_inv_idempotent:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3408
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3409
    and a1: "a$1 \<noteq> 0"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3410
  shows "fps_inv (fps_inv a) = a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3411
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3412
  let ?r = "fps_inv"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3413
  have ra0: "?r a $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3414
    by (simp add: fps_inv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3415
  from a1 have ra1: "?r a $ 1 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3416
    by (simp add: fps_inv_def field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3417
  have X0: "X$0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3418
    by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3419
  from fps_inv[OF ra0 ra1] have "?r (?r a) oo ?r a = X" .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3420
  then have "?r (?r a) oo ?r a oo a = X oo a"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3421
    by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3422
  then have "?r (?r a) oo (?r a oo a) = a"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3423
    unfolding X_fps_compose_startby0[OF a0]
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3424
    unfolding fps_compose_assoc[OF a0 ra0, symmetric] .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3425
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3426
    unfolding fps_inv[OF a0 a1] by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3427
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3428
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3429
lemma fps_ginv_ginv:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3430
  assumes a0: "a$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3431
    and a1: "a$1 \<noteq> 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3432
    and c0: "c$0 = 0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3433
    and  c1: "c$1 \<noteq> 0"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3434
  shows "fps_ginv b (fps_ginv c a) = b oo a oo fps_inv c"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3435
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3436
  let ?r = "fps_ginv"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3437
  from c0 have rca0: "?r c a $0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3438
    by (simp add: fps_ginv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3439
  from a1 c1 have rca1: "?r c a $ 1 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3440
    by (simp add: fps_ginv_def field_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3441
  from fps_ginv[OF rca0 rca1]
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3442
  have "?r b (?r c a) oo ?r c a = b" .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3443
  then have "?r b (?r c a) oo ?r c a oo a = b oo a"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3444
    by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3445
  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
  3446
    apply (subst fps_compose_assoc)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3447
    using a0 c0
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3448
    apply (auto simp add: fps_ginv_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3449
    done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3450
  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
  3451
    unfolding fps_ginv[OF a0 a1] .
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3452
  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
  3453
    by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3454
  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
  3455
    apply (subst fps_compose_assoc)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3456
    using a0 c0
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3457
    apply (auto simp add: fps_inv_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3458
    done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3459
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3460
    unfolding fps_inv_right[OF c0 c1] by simp
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3461
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3462
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3463
lemma fps_ginv_deriv:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3464
  assumes a0:"a$0 = (0::'a::field)"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3465
    and a1: "a$1 \<noteq> 0"
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3466
  shows "fps_deriv (fps_ginv b a) = (fps_deriv b / fps_deriv a) oo fps_ginv X a"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3467
proof -
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3468
  let ?ia = "fps_ginv b a"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3469
  let ?iXa = "fps_ginv X a"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3470
  let ?d = "fps_deriv"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3471
  let ?dia = "?d ?ia"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3472
  have iXa0: "?iXa $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3473
    by (simp add: fps_ginv_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3474
  have da0: "?d a $ 0 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3475
    using a1 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3476
  from fps_ginv[OF a0 a1, of b] have "?d (?ia oo a) = fps_deriv b"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3477
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3478
  then have "(?d ?ia oo a) * ?d a = ?d b"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3479
    unfolding fps_compose_deriv[OF a0] .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3480
  then have "(?d ?ia oo a) * ?d a * inverse (?d a) = ?d b * inverse (?d a)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3481
    by simp
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3482
  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
  3483
    by (simp add: fps_divide_unit)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3484
  then have "(?d ?ia oo a) oo ?iXa =  (?d b / ?d a) oo ?iXa"
32410
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3485
    unfolding inverse_mult_eq_1[OF da0] by simp
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3486
  then have "?d ?ia oo (a oo ?iXa) =  (?d b / ?d a) oo ?iXa"
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3487
    unfolding fps_compose_assoc[OF iXa0 a0] .
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3488
  then show ?thesis unfolding fps_inv_ginv[symmetric]
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3489
    unfolding fps_inv_right[OF a0 a1] by simp
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3490
qed
624bd2ea7c1e Derivative of general reverses
chaieb
parents: 31075
diff changeset
  3491
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3492
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3493
subsection \<open>Elementary series\<close>
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3494
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3495
subsubsection \<open>Exponential series\<close>
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3496
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3497
definition "E x = Abs_fps (\<lambda>n. x^n / of_nat (fact n))"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3498
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3499
lemma E_deriv[simp]: "fps_deriv (E a) = fps_const (a::'a::field_char_0) * E a" (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3500
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3501
  have "?l$n = ?r $ n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3502
    apply (auto simp add: E_def field_simps power_Suc[symmetric]
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3503
      simp del: fact.simps of_nat_Suc power_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3504
    apply (simp add: of_nat_mult field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3505
    done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3506
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3507
    by (simp add: fps_eq_iff)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3508
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3509
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3510
lemma E_unique_ODE:
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3511
  "fps_deriv a = fps_const c * a \<longleftrightarrow> a = fps_const (a$0) * E (c::'a::field_char_0)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3512
  (is "?lhs \<longleftrightarrow> ?rhs")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3513
proof
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3514
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3515
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3516
    from that have th: "\<And>n. a $ Suc n = c * a$n / of_nat (Suc n)"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3517
      by (simp add: fps_deriv_def fps_eq_iff field_simps del: of_nat_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3518
    have th': "a$n = a$0 * c ^ n/ (fact n)" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3519
    proof (induct n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3520
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3521
      then show ?case by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3522
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3523
      case Suc
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3524
      then show ?case
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3525
        unfolding th
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3526
        using fact_gt_zero
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3527
        apply (simp add: field_simps del: of_nat_Suc fact_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3528
        apply simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3529
        done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3530
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3531
    show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3532
      by (auto simp add: fps_eq_iff fps_const_mult_left E_def intro: th')
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3533
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3534
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3535
    using that by (metis E_deriv fps_deriv_mult_const_left mult.left_commute)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3536
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3537
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3538
lemma E_add_mult: "E (a + b) = E (a::'a::field_char_0) * E b" (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3539
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3540
  have "fps_deriv ?r = fps_const (a + b) * ?r"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3541
    by (simp add: fps_const_add[symmetric] field_simps del: fps_const_add)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3542
  then have "?r = ?l"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3543
    by (simp only: E_unique_ODE) (simp add: fps_mult_nth E_def)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3544
  then show ?thesis ..
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3545
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3546
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3547
lemma E_nth[simp]: "E a $ n = a^n / of_nat (fact n)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3548
  by (simp add: E_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3549
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3550
lemma E0[simp]: "E (0::'a::field) = 1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3551
  by (simp add: fps_eq_iff power_0_left)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3552
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3553
lemma E_neg: "E (- a) = inverse (E (a::'a::field_char_0))"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3554
proof -
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3555
  from E_add_mult[of a "- a"] have th0: "E a * E (- a) = 1" by simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3556
  from fps_inverse_unique[OF th0] show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3557
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3558
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3559
lemma E_nth_deriv[simp]: "fps_nth_deriv n (E (a::'a::field_char_0)) = (fps_const a)^n * (E a)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3560
  by (induct n) auto
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3561
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3562
lemma X_compose_E[simp]: "X oo E (a::'a::field) = E a - 1"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3563
  by (simp add: fps_eq_iff X_fps_compose)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3564
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3565
lemma LE_compose:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3566
  assumes a: "a \<noteq> 0"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3567
  shows "fps_inv (E a - 1) oo (E a - 1) = X"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3568
    and "(E a - 1) oo fps_inv (E a - 1) = X"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3569
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3570
  let ?b = "E a - 1"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3571
  have b0: "?b $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3572
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3573
  have b1: "?b $ 1 \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3574
    by (simp add: a)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3575
  from fps_inv[OF b0 b1] show "fps_inv (E a - 1) oo (E a - 1) = X" .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3576
  from fps_inv_right[OF b0 b1] show "(E a - 1) oo fps_inv (E a - 1) = X" .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3577
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3578
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3579
lemma E_power_mult: "(E (c::'a::field_char_0))^n = E (of_nat n * c)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3580
  by (induct n) (auto simp add: field_simps E_add_mult)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3581
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3582
lemma radical_E:
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3583
  assumes r: "r (Suc k) 1 = 1"
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3584
  shows "fps_radical r (Suc k) (E (c::'a::field_char_0)) = E (c / of_nat (Suc k))"
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3585
proof -
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3586
  let ?ck = "(c / of_nat (Suc k))"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3587
  let ?r = "fps_radical r (Suc k)"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3588
  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
  3589
    by (simp_all del: of_nat_Suc)
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3590
  have th0: "E ?ck ^ (Suc k) = E c" unfolding E_power_mult eq0 ..
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3591
  have th: "r (Suc k) (E c $0) ^ Suc k = E c $ 0"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3592
    "r (Suc k) (E c $ 0) = E ?ck $ 0" "E c $ 0 \<noteq> 0" using r by simp_all
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3593
  from th0 radical_unique[where r=r and k=k, OF th] show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3594
    by auto
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3595
qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3596
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  3597
lemma Ec_E1_eq: "E (1::'a::field_char_0) oo (fps_const c * X) = E c"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3598
  apply (auto simp add: fps_eq_iff E_def fps_compose_def power_mult_distrib)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3599
  apply (simp add: cond_value_iff cond_application_beta setsum.delta' cong del: if_weak_cong)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3600
  done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3601
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3602
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3603
subsubsection \<open>Logarithmic series\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3604
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3605
lemma Abs_fps_if_0:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3606
  "Abs_fps (\<lambda>n. if n = 0 then (v::'a::ring_1) else f n) =
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3607
    fps_const v + X * Abs_fps (\<lambda>n. f (Suc n))"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3608
  by (auto simp add: fps_eq_iff)
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3609
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3610
definition L :: "'a::field_char_0 \<Rightarrow> 'a fps"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  3611
  where "L c = fps_const (1/c) * Abs_fps (\<lambda>n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n)"
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3612
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3613
lemma fps_deriv_L: "fps_deriv (L c) = fps_const (1/c) * inverse (1 + X)"
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3614
  unfolding fps_inverse_X_plus1
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3615
  by (simp add: L_def fps_eq_iff del: of_nat_Suc)
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3616
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3617
lemma L_nth: "L c $ n = (if n = 0 then 0 else 1/c * ((- 1) ^ (n - 1) / of_nat n))"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3618
  by (simp add: L_def field_simps)
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3619
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3620
lemma L_0[simp]: "L c $ 0 = 0" by (simp add: L_def)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3621
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3622
lemma L_E_inv:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3623
  fixes a :: "'a::field_char_0"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3624
  assumes a: "a \<noteq> 0"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3625
  shows "L a = fps_inv (E a - 1)"  (is "?l = ?r")
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3626
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3627
  let ?b = "E a - 1"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3628
  have b0: "?b $ 0 = 0" by simp
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3629
  have b1: "?b $ 1 \<noteq> 0" by (simp add: a)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3630
  have "fps_deriv (E a - 1) oo fps_inv (E a - 1) =
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3631
    (fps_const a * (E a - 1) + fps_const a) oo fps_inv (E a - 1)"
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3632
    by (simp add: field_simps)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3633
  also have "\<dots> = fps_const a * (X + 1)"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3634
    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
  3635
    apply (simp add: field_simps)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3636
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3637
  finally have eq: "fps_deriv (E a - 1) oo fps_inv (E a - 1) = fps_const a * (X + 1)" .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3638
  from fps_inv_deriv[OF b0 b1, unfolded eq]
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3639
  have "fps_deriv (fps_inv ?b) = fps_const (inverse a) / (X + 1)"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3640
    using a
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3641
    by (simp add: fps_const_inverse eq fps_divide_def fps_inverse_mult)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3642
  then have "fps_deriv ?l = fps_deriv ?r"
57512
cc97b347b301 reduced name variants for assoc and commute on plus and mult
haftmann
parents: 57418
diff changeset
  3643
    by (simp add: fps_deriv_L add.commute fps_divide_def divide_inverse)
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3644
  then show ?thesis unfolding fps_deriv_eq_iff
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3645
    by (simp add: L_nth fps_inv_def)
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3646
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3647
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3648
lemma L_mult_add:
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3649
  assumes c0: "c\<noteq>0"
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3650
    and d0: "d\<noteq>0"
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3651
  shows "L c + L d = fps_const (c+d) * L (c*d)"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3652
  (is "?r = ?l")
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3653
proof-
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3654
  from c0 d0 have eq: "1/c + 1/d = (c+d)/(c*d)" by (simp add: field_simps)
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3655
  have "fps_deriv ?r = fps_const (1/c + 1/d) * inverse (1 + X)"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3656
    by (simp add: fps_deriv_L fps_const_add[symmetric] algebra_simps del: fps_const_add)
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3657
  also have "\<dots> = fps_deriv ?l"
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3658
    apply (simp add: fps_deriv_L)
52903
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3659
    apply (simp add: fps_eq_iff eq)
6c89225ddeba tuned proofs;
wenzelm
parents: 52902
diff changeset
  3660
    done
31369
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3661
  finally show ?thesis
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3662
    unfolding fps_deriv_eq_iff by simp
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3663
qed
8b460fd12100 Reverses idempotent; radical of E; generalized logarithm;
chaieb
parents: 31199
diff changeset
  3664
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3665
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3666
subsubsection \<open>Binomial series\<close>
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3667
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3668
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
  3669
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3670
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
  3671
  by (simp add: fps_binomial_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3672
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3673
lemma fps_binomial_ODE_unique:
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3674
  fixes c :: "'a::field_char_0"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3675
  shows "fps_deriv a = (fps_const c * a) / (1 + X) \<longleftrightarrow> a = fps_const (a$0) * fps_binomial c"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3676
  (is "?lhs \<longleftrightarrow> ?rhs")
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3677
proof
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3678
  let ?da = "fps_deriv a"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3679
  let ?x1 = "(1 + X):: 'a fps"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3680
  let ?l = "?x1 * ?da"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3681
  let ?r = "fps_const c * a"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3682
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3683
  have eq: "?l = ?r \<longleftrightarrow> ?lhs"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3684
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3685
    have x10: "?x1 $ 0 \<noteq> 0" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3686
    have "?l = ?r \<longleftrightarrow> inverse ?x1 * ?l = inverse ?x1 * ?r" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3687
    also have "\<dots> \<longleftrightarrow> ?da = (fps_const c * a) / ?x1"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3688
      apply (simp only: fps_divide_def  mult.assoc[symmetric] inverse_mult_eq_1[OF x10])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3689
      apply (simp add: field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3690
      done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3691
    finally show ?thesis .
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3692
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3693
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3694
  show ?rhs if ?lhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3695
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3696
    from eq that have h: "?l = ?r" ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3697
    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
  3698
    proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3699
      from h have "?l $ n = ?r $ n" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3700
      then show ?thesis
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3701
        apply (simp add: field_simps del: of_nat_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3702
        apply (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3703
        apply (simp_all add: field_simps del: of_nat_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3704
        done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3705
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3706
    have th1: "a $ n = (c gchoose n) * a $ 0" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3707
    proof (induct n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3708
      case 0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3709
      then show ?case by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3710
    next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3711
      case (Suc m)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3712
      then show ?case
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3713
        unfolding th0
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3714
        apply (simp add: field_simps del: of_nat_Suc)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3715
        unfolding mult.assoc[symmetric] gbinomial_mult_1
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3716
        apply (simp add: field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3717
        done
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3718
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3719
    show ?thesis
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3720
      apply (simp add: fps_eq_iff)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3721
      apply (subst th1)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3722
      apply (simp add: field_simps)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3723
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3724
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3725
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3726
  show ?lhs if ?rhs
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3727
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3728
    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
  3729
      by (simp add: mult.commute)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3730
    have "?l = ?r"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3731
      apply (subst \<open>?rhs\<close>)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3732
      apply (subst (2) \<open>?rhs\<close>)
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3733
      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
  3734
      unfolding mult.assoc[symmetric] th00 gbinomial_mult_1
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3735
      apply (simp add: field_simps gbinomial_mult_1)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3736
      done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3737
    with eq show ?thesis ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3738
  qed
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3739
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3740
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3741
lemma fps_binomial_deriv: "fps_deriv (fps_binomial c) = fps_const c * fps_binomial c / (1 + X)"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3742
proof -
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3743
  let ?a = "fps_binomial c"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3744
  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
  3745
  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
  3746
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3747
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3748
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
  3749
proof -
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3750
  let ?P = "?r - ?l"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3751
  let ?b = "fps_binomial"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3752
  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
  3753
  have "fps_deriv ?P = ?db c * ?b d + ?b c * ?db d - ?db (c + d)"  by simp
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3754
  also have "\<dots> = inverse (1 + X) *
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3755
      (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
  3756
    unfolding fps_binomial_deriv
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3757
    by (simp add: fps_divide_def field_simps)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3758
  also have "\<dots> = (fps_const (c + d)/ (1 + X)) * ?P"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  3759
    by (simp add: field_simps fps_divide_unit fps_const_add[symmetric] del: fps_const_add)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3760
  finally have th0: "fps_deriv ?P = fps_const (c+d) * ?P / (1 + X)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3761
    by (simp add: fps_divide_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3762
  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
  3763
    unfolding fps_binomial_ODE_unique[symmetric]
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3764
    using th0 by simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3765
  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
  3766
  then show ?thesis by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3767
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3768
61552
980dd46a03fb Added binomial identities to CONTRIBUTORS; small lemmas on of_int/pochhammer
eberlm
parents: 60867
diff changeset
  3769
lemma fps_binomial_minus_one: "fps_binomial (- 1) = inverse (1 + X)"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3770
  (is "?l = inverse ?r")
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3771
proof-
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3772
  have th: "?r$0 \<noteq> 0" by simp
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3773
  have th': "fps_deriv (inverse ?r) = fps_const (- 1) * inverse ?r / (1 + X)"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3774
    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
  3775
      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
  3776
  have eq: "inverse ?r $ 0 = 1"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3777
    by (simp add: fps_inverse_def)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3778
  from iffD1[OF fps_binomial_ODE_unique[of "inverse (1 + X)" "- 1"] th'] eq
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3779
  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
  3780
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3781
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3782
text \<open>Vandermonde's Identity as a consequence.\<close>
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3783
lemma gbinomial_Vandermonde:
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3784
  "setsum (\<lambda>k. (a gchoose k) * (b gchoose (n - k))) {0..n} = (a + b) gchoose n"
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3785
proof -
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3786
  let ?ba = "fps_binomial a"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3787
  let ?bb = "fps_binomial b"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3788
  let ?bab = "fps_binomial (a + b)"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3789
  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
  3790
  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
  3791
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3792
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3793
lemma binomial_Vandermonde:
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3794
  "setsum (\<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
  3795
  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
  3796
  by (simp only: binomial_gbinomial[symmetric] of_nat_mult[symmetric]
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
  3797
                 of_nat_setsum[symmetric] of_nat_add[symmetric] of_nat_eq_iff)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3798
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3799
lemma binomial_Vandermonde_same: "setsum (\<lambda>k. (n choose k)\<^sup>2) {0..n} = (2 * n) choose n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3800
  using binomial_Vandermonde[of n n n, symmetric]
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3801
  unfolding mult_2
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3802
  apply (simp add: power2_eq_square)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3803
  apply (rule setsum.cong)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3804
  apply (auto intro:  binomial_symmetric)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3805
  done
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3806
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3807
lemma Vandermonde_pochhammer_lemma:
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3808
  fixes a :: "'a::field_char_0"
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3809
  assumes b: "\<forall>j\<in>{0 ..<n}. b \<noteq> of_nat j"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3810
  shows "setsum (\<lambda>k. (pochhammer (- a) k * pochhammer (- (of_nat n)) k) /
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3811
      (of_nat (fact k) * pochhammer (b - of_nat n + 1) k)) {0..n} =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3812
    pochhammer (- (a + b)) n / pochhammer (- b) n"
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3813
  (is "?l = ?r")
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3814
proof -
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3815
  let ?m1 = "\<lambda>m. (- 1 :: 'a) ^ m"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3816
  let ?f = "\<lambda>m. of_nat (fact m)"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3817
  let ?p = "\<lambda>(x::'a). pochhammer (- x)"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3818
  from b have bn0: "?p b n \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3819
    unfolding pochhammer_eq_0_iff by simp
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3820
  have th00:
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3821
    "b gchoose (n - k) =
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3822
        (?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
  3823
      (is ?gchoose)
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3824
    "pochhammer (1 + b - of_nat n) k \<noteq> 0"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3825
      (is ?pochhammer)
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3826
    if kn: "k \<in> {0..n}" for k
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3827
  proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3828
    have nz: "pochhammer (1 + b - of_nat n) n \<noteq> 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3829
    proof
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3830
      assume "pochhammer (1 + b - of_nat n) n = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3831
      then have c: "pochhammer (b - of_nat n + 1) n = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3832
        by (simp add: algebra_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3833
      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
  3834
        unfolding pochhammer_eq_0_iff by blast
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3835
      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
  3836
        by (simp add: algebra_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3837
      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
  3838
        using j kn by (simp add: of_nat_diff)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3839
      with b show False using j by auto
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3840
    qed
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3841
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3842
    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
  3843
      by (rule pochhammer_neq_0_mono)
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3844
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  3845
    consider "k = 0 \<or> n = 0" | "k \<noteq> 0" "n \<noteq> 0"
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  3846
      by blast
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3847
    then have "b gchoose (n - k) =
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3848
      (?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
  3849
    proof cases
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3850
      case 1
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3851
      then show ?thesis
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3852
        using kn by (cases "k = 0") (simp_all add: gbinomial_pochhammer)
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3853
    next
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  3854
      case neq: 2
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3855
      then obtain m where m: "n = Suc m"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3856
        by (cases n) auto
60567
19c277ea65ae tuned proofs -- less digits;
wenzelm
parents: 60558
diff changeset
  3857
      from neq(1) obtain h where h: "k = Suc h"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3858
        by (cases k) auto
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3859
      show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3860
      proof (cases "k = n")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3861
        case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3862
        then show ?thesis
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59815
diff changeset
  3863
          using pochhammer_minus'[where k=k and b=b]
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59815
diff changeset
  3864
          apply (simp add: pochhammer_same)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3865
          using bn0
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3866
          apply (simp add: field_simps power_add[symmetric])
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3867
          done
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3868
      next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3869
        case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3870
        with kn have kn': "k < n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3871
          by simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3872
        have m1nk: "?m1 n = setprod (\<lambda>i. - 1) {0..m}" "?m1 k = setprod (\<lambda>i. - 1) {0..h}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3873
          by (simp_all add: setprod_constant m h)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3874
        have bnz0: "pochhammer (b - of_nat n + 1) k \<noteq> 0"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3875
          using bn0 kn
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3876
          unfolding pochhammer_eq_0_iff
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3877
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3878
          apply (erule_tac x= "n - ka - 1" in allE)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3879
          apply (auto simp add: algebra_simps of_nat_diff)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3880
          done
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3881
        have eq1: "setprod (\<lambda>k. (1::'a) + of_nat m - of_nat k) {0 .. h} =
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3882
          setprod 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
  3883
          using kn' h m
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  3884
          by (intro setprod.reindex_bij_witness[where i="\<lambda>k. Suc m - k" and j="\<lambda>k. Suc m - k"])
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  3885
             (auto simp: of_nat_diff)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3886
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3887
        have th1: "(?m1 k * ?p (of_nat n) k) / ?f n = 1 / of_nat(fact (n - k))"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3888
          unfolding m1nk
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3889
          unfolding m h pochhammer_Suc_setprod
54489
03ff4d1e6784 eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents: 54452
diff changeset
  3890
          apply (simp add: field_simps del: fact_Suc)
59730
b7c394c7a619 The factorial function, "fact", now has type "nat => 'a"
paulson <lp15@cam.ac.uk>
parents: 59667
diff changeset
  3891
          unfolding fact_altdef id_def
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3892
          unfolding of_nat_setprod
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3893
          unfolding setprod.distrib[symmetric]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3894
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3895
          unfolding eq1
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3896
          apply (subst setprod.union_disjoint[symmetric])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3897
          apply (auto)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3898
          apply (rule setprod.cong)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3899
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3900
          done
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3901
        have th20: "?m1 n * ?p b n = setprod (\<lambda>i. b - of_nat i) {0..m}"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3902
          unfolding m1nk
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3903
          unfolding m h pochhammer_Suc_setprod
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3904
          unfolding setprod.distrib[symmetric]
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3905
          apply (rule setprod.cong)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3906
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3907
          done
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3908
        have th21:"pochhammer (b - of_nat n + 1) k = setprod (\<lambda>i. b - of_nat i) {n - k .. n - 1}"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3909
          unfolding h m
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3910
          unfolding pochhammer_Suc_setprod
57129
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  3911
          using kn m h
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  3912
          by (intro setprod.reindex_bij_witness[where i="\<lambda>k. n - 1 - k" and j="\<lambda>i. m-i"])
7edb7550663e introduce more powerful reindexing rules for big operators
hoelzl
parents: 56480
diff changeset
  3913
             (auto simp: of_nat_diff)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3914
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3915
        have "?m1 n * ?p b n =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3916
          pochhammer (b - of_nat n + 1) k * setprod (\<lambda>i. b - of_nat i) {0.. n - k - 1}"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3917
          unfolding th20 th21
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3918
          unfolding h m
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3919
          apply (subst setprod.union_disjoint[symmetric])
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3920
          using kn' h m
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3921
          apply auto
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3922
          apply (rule setprod.cong)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3923
          apply auto
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3924
          done
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3925
        then have th2: "(?m1 n * ?p b n)/pochhammer (b - of_nat n + 1) k =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3926
          setprod (\<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
  3927
          using nz' by (simp add: field_simps)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3928
        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
  3929
          ((?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
  3930
          using bnz0
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3931
          by (simp add: field_simps)
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3932
        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
  3933
          unfolding th1 th2
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32456
diff changeset
  3934
          using kn' by (simp add: gbinomial_def)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3935
        finally show ?thesis by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3936
      qed
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3937
    qed
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3938
    then show ?gchoose and ?pochhammer
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3939
      apply (cases "n = 0")
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3940
      using nz'
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3941
      apply auto
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3942
      done
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3943
  qed
60504
17741c71a731 tuned proofs;
wenzelm
parents: 60501
diff changeset
  3944
  have "?r = ((a + b) gchoose n) * (of_nat (fact n) / (?m1 n * pochhammer (- b) n))"
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3945
    unfolding gbinomial_pochhammer
36350
bc7982c54e37 dropped group_simps, ring_simps, field_eq_simps
haftmann
parents: 36311
diff changeset
  3946
    using bn0 by (auto simp add: field_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3947
  also have "\<dots> = ?l"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3948
    unfolding gbinomial_Vandermonde[symmetric]
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3949
    apply (simp add: th00)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3950
    unfolding gbinomial_pochhammer
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3951
    using bn0
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3952
    apply (simp add: setsum_left_distrib setsum_right_distrib field_simps)
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3953
    apply (rule setsum.cong)
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  3954
    apply (rule refl)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3955
    apply (drule th00(2))
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3956
    apply (simp add: field_simps power_add[symmetric])
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  3957
    done
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3958
  finally show ?thesis by simp
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3959
qed
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  3960
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3961
lemma Vandermonde_pochhammer:
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3962
  fixes a :: "'a::field_char_0"
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3963
  assumes c: "\<forall>i \<in> {0..< n}. c \<noteq> - of_nat i"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  3964
  shows "setsum (\<lambda>k. (pochhammer a k * pochhammer (- (of_nat n)) k) /
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3965
    (of_nat (fact k) * pochhammer c k)) {0..n} = pochhammer (c - a) n / pochhammer c n"
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3966
proof -
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3967
  let ?a = "- a"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3968
  let ?b = "c + of_nat n - 1"
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3969
  have h: "\<forall> j \<in>{0..< n}. ?b \<noteq> of_nat j"
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  3970
    using c
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3971
    apply (auto simp add: algebra_simps of_nat_diff)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3972
    apply (erule_tac x = "n - j - 1" in ballE)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3973
    apply (auto simp add: of_nat_diff algebra_simps)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3974
    done
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3975
  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
  3976
    unfolding pochhammer_minus
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3977
    by (simp add: algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3978
  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
  3979
    unfolding pochhammer_minus
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3980
    by simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3981
  have nz: "pochhammer c n \<noteq> 0" using c
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3982
    by (simp add: pochhammer_eq_0_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3983
  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
  3984
  show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3985
    using nz by (simp add: field_simps setsum_right_distrib)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  3986
qed
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  3987
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  3988
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  3989
subsubsection \<open>Formal trigonometric functions\<close>
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3990
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3991
definition "fps_sin (c::'a::field_char_0) =
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3992
  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
  3993
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3994
definition "fps_cos (c::'a::field_char_0) =
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  3995
  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
  3996
30488
5c4c3a9e9102 remove trailing spaces
huffman
parents: 30273
diff changeset
  3997
lemma fps_sin_deriv:
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3998
  "fps_deriv (fps_sin c) = fps_const c * fps_cos c"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  3999
  (is "?lhs = ?rhs")
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  4000
proof (rule fps_ext)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4001
  fix n :: nat
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4002
  show "?lhs $ n = ?rhs $ n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4003
  proof (cases "even n")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4004
    case True
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4005
    have "?lhs$n = of_nat (n+1) * (fps_sin c $ (n+1))" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4006
    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
  4007
      using True by (simp add: fps_sin_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4008
    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
  4009
      unfolding fact_Suc of_nat_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4010
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4011
    also have "\<dots> = (- 1)^(n div 2) *c^Suc n / of_nat (fact n)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4012
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4013
    finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4014
      using True by (simp add: fps_cos_def field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4015
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4016
    case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4017
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4018
      by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4019
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4020
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4021
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4022
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
  4023
  (is "?lhs = ?rhs")
31273
da95bc889ad2 use class field_char_0 for fps definitions
huffman
parents: 31199
diff changeset
  4024
proof (rule fps_ext)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4025
  have th0: "- ((- 1::'a) ^ n) = (- 1)^Suc n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4026
    by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4027
  show "?lhs $ n = ?rhs $ n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4028
  proof (cases "even n")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4029
    case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4030
    then have n0: "n \<noteq> 0" by presburger
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4031
    from False have th1: "Suc ((n - 1) div 2) = Suc n div 2"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4032
      by (cases n) simp_all
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4033
    have "?lhs$n = of_nat (n+1) * (fps_cos c $ (n+1))" by simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4034
    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
  4035
      using False by (simp add: fps_cos_def)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4036
    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
  4037
      unfolding fact_Suc of_nat_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4038
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4039
    also have "\<dots> = (- 1)^((n + 1) div 2) * c^Suc n / of_nat (fact n)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4040
      by (simp add: field_simps del: of_nat_add of_nat_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4041
    also have "\<dots> = (- ((- 1)^((n - 1) div 2))) * c^Suc n / of_nat (fact n)"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4042
      unfolding th0 unfolding th1 by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4043
    finally show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4044
      using False by (simp add: fps_sin_def field_simps)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4045
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4046
    case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4047
    then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4048
      by (simp_all add: fps_deriv_def fps_sin_def fps_cos_def)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4049
  qed
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4050
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4051
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4052
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
  4053
  (is "?lhs = _")
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  4054
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4055
  have "fps_deriv ?lhs = 0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4056
    apply (simp add:  fps_deriv_power fps_sin_deriv fps_cos_deriv)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4057
    apply (simp add: field_simps fps_const_neg[symmetric] del: fps_const_neg)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4058
    done
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4059
  then have "?lhs = fps_const (?lhs $ 0)"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4060
    unfolding fps_deriv_eq_0_iff .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4061
  also have "\<dots> = 1"
30960
fec1a04b7220 power operation defined generic
haftmann
parents: 30952
diff changeset
  4062
    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
  4063
  finally show ?thesis .
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4064
qed
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4065
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4066
lemma fps_sin_nth_0 [simp]: "fps_sin c $ 0 = 0"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4067
  unfolding fps_sin_def by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4068
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4069
lemma fps_sin_nth_1 [simp]: "fps_sin c $ 1 = c"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4070
  unfolding fps_sin_def by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4071
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4072
lemma fps_sin_nth_add_2:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4073
    "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
  4074
  unfolding fps_sin_def
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4075
  apply (cases n)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4076
  apply simp
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4077
  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc)
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4078
  apply (simp add: of_nat_mult del: of_nat_Suc mult_Suc)
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4079
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4080
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4081
lemma fps_cos_nth_0 [simp]: "fps_cos c $ 0 = 1"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4082
  unfolding fps_cos_def by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4083
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4084
lemma fps_cos_nth_1 [simp]: "fps_cos c $ 1 = 0"
53195
e4b18828a817 tuned proofs;
wenzelm
parents: 53077
diff changeset
  4085
  unfolding fps_cos_def by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4086
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4087
lemma fps_cos_nth_add_2:
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4088
  "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
  4089
  unfolding fps_cos_def
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4090
  apply (simp add: nonzero_divide_eq_eq nonzero_eq_divide_eq del: of_nat_Suc fact_Suc)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4091
  apply (simp add: of_nat_mult del: of_nat_Suc mult_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4092
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4093
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4094
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
  4095
  unfolding One_nat_def numeral_2_eq_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4096
  apply (induct n rule: nat_less_induct)
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4097
  apply (case_tac n)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4098
  apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4099
  apply (rename_tac m)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4100
  apply (case_tac m)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4101
  apply simp
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4102
  apply (rename_tac k)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4103
  apply (case_tac k)
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4104
  apply simp_all
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4105
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4106
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4107
lemma nat_add_1_add_1: "(n::nat) + 1 + 1 = n + 2"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4108
  by simp
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4109
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4110
lemma eq_fps_sin:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4111
  assumes 0: "a $ 0 = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4112
    and 1: "a $ 1 = c"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4113
    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
  4114
  shows "a = fps_sin c"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4115
  apply (rule fps_ext)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4116
  apply (induct_tac n rule: nat_induct2)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4117
  apply (simp add: 0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4118
  apply (simp add: 1 del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4119
  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4120
  apply (simp add: nat_add_1_add_1 fps_sin_nth_add_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4121
              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4122
  apply (subst minus_divide_left)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4123
  apply (subst nonzero_eq_divide_eq)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4124
  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
  4125
  apply (simp only: ac_simps)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4126
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4127
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4128
lemma eq_fps_cos:
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4129
  assumes 0: "a $ 0 = 1"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4130
    and 1: "a $ 1 = 0"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4131
    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
  4132
  shows "a = fps_cos c"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4133
  apply (rule fps_ext)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4134
  apply (induct_tac n rule: nat_induct2)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4135
  apply (simp add: 0)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4136
  apply (simp add: 1 del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4137
  apply (rename_tac m, cut_tac f="\<lambda>a. a $ m" in arg_cong [OF 2])
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4138
  apply (simp add: nat_add_1_add_1 fps_cos_nth_add_2
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4139
              del: One_nat_def of_nat_Suc of_nat_add add_2_eq_Suc')
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4140
  apply (subst minus_divide_left)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4141
  apply (subst nonzero_eq_divide_eq)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4142
  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
  4143
  apply (simp only: ac_simps)
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4144
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4145
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4146
lemma mult_nth_0 [simp]: "(a * b) $ 0 = a $ 0 * b $ 0"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4147
  by (simp add: fps_mult_nth)
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4148
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4149
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
  4150
  by (simp add: fps_mult_nth)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4151
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4152
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
  4153
  apply (rule eq_fps_sin [symmetric], simp, simp del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4154
  apply (simp del: fps_const_neg fps_const_add fps_const_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4155
              add: fps_const_add [symmetric] fps_const_neg [symmetric]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4156
                   fps_sin_deriv fps_cos_deriv algebra_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4157
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4158
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4159
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
  4160
  apply (rule eq_fps_cos [symmetric], simp, simp del: One_nat_def)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4161
  apply (simp del: fps_const_neg fps_const_add fps_const_mult
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4162
              add: fps_const_add [symmetric] fps_const_neg [symmetric]
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4163
                   fps_sin_deriv fps_cos_deriv algebra_simps)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4164
  done
31274
d2b5c6b07988 addition formulas for fps_sin, fps_cos
huffman
parents: 31273
diff changeset
  4165
31968
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  4166
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
  4167
  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
  4168
0314441a53a6 FPS form a metric space, which justifies the infinte sum notation
chaieb
parents: 31790
diff changeset
  4169
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
  4170
  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
  4171
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4172
definition "fps_tan c = fps_sin c / fps_cos c"
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4173
53077
a1b3784f8129 more symbols;
wenzelm
parents: 52903
diff changeset
  4174
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
  4175
proof -
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4176
  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
  4177
  from this have "fps_cos c \<noteq> 0" by (intro notI) simp
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4178
  hence "fps_deriv (fps_tan c) = 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4179
           fps_const c * (fps_cos c^2 + fps_sin c^2) / (fps_cos c^2)"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4180
    by (simp add: fps_tan_def fps_divide_deriv power2_eq_square algebra_simps 
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4181
                  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
  4182
             del: fps_const_neg)
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4183
  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
  4184
  finally show ?thesis by simp
29687
4d934a895d11 A formalization of formal power series
chaieb
parents:
diff changeset
  4185
qed
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  4186
60558
4fcc6861e64f tuned proofs;
wenzelm
parents: 60504
diff changeset
  4187
text \<open>Connection to E c over the complex numbers --- Euler and de Moivre.\<close>
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4188
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4189
lemma Eii_sin_cos: "E (ii * c) = fps_cos c + fps_const ii * fps_sin c"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4190
  (is "?l = ?r")
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4191
proof -
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4192
  have "?l $ n = ?r $ n" for n
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4193
  proof (cases "even n")
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4194
    case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4195
    then obtain m where m: "n = 2 * m" ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4196
    show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4197
      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
  4198
  next
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4199
    case False
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4200
    then obtain m where m: "n = 2 * m + 1" ..
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4201
    show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4202
      by (simp add: m fps_sin_def fps_cos_def power_mult_distrib
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4203
        power_mult power_minus [of "c ^ 2"])
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4204
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4205
  then show ?thesis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4206
    by (simp add: fps_eq_iff)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4207
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4208
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4209
lemma E_minus_ii_sin_cos: "E (- (ii * c)) = fps_cos c - fps_const ii * fps_sin c"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4210
  unfolding minus_mult_right Eii_sin_cos by (simp add: fps_sin_even fps_cos_odd)
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4211
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4212
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
  4213
  by (fact fps_const_sub)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4214
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  4215
lemma fps_numeral_fps_const: "numeral i = fps_const (numeral i :: 'a::comm_ring_1)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  4216
  by (fact numeral_fps_const) (* FIXME: duplicate *)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4217
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4218
lemma fps_cos_Eii: "fps_cos c = (E (ii * c) + E (- ii * c)) / fps_const 2"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4219
proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4220
  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
  4221
    by (simp add: numeral_fps_const)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4222
  show ?thesis
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4223
    unfolding Eii_sin_cos minus_mult_commute
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4224
    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
  4225
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4226
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4227
lemma fps_sin_Eii: "fps_sin c = (E (ii * c) - E (- ii * c)) / fps_const (2*ii)"
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4228
proof -
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4229
  have th: "fps_const \<i> * fps_sin c + fps_const \<i> * fps_sin c = fps_sin c * fps_const (2 * ii)"
47108
2a1953f0d20d merged fork with new numeral representation (see NEWS)
huffman
parents: 46757
diff changeset
  4230
    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
  4231
  show ?thesis
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4232
    unfolding Eii_sin_cos minus_mult_commute
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4233
    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
  4234
qed
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4235
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4236
lemma fps_tan_Eii:
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4237
  "fps_tan c = (E (ii * c) - E (- ii * c)) / (fps_const ii * (E (ii * c) + E (- ii * c)))"
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4238
  unfolding fps_tan_def fps_sin_Eii fps_cos_Eii mult_minus_left E_neg
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4239
  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
  4240
  apply simp
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4241
  done
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4242
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4243
lemma fps_demoivre:
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4244
  "(fps_cos a + fps_const ii * fps_sin a)^n =
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4245
    fps_cos (of_nat n * a) + fps_const ii * fps_sin (of_nat n * a)"
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4246
  unfolding Eii_sin_cos[symmetric] E_power_mult
57514
bdc2c6b40bf2 prefer ac_simps collections over separate name bindings for add and mult
haftmann
parents: 57512
diff changeset
  4247
  by (simp add: ac_simps)
32157
adea7a729c7a Moved important theorems from FPS_Examples to FPS --- they are not
chaieb
parents: 31968
diff changeset
  4248
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4249
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  4250
subsection \<open>Hypergeometric series\<close>
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4251
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59862
diff changeset
  4252
definition "F as bs (c::'a::{field_char_0,field}) =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4253
  Abs_fps (\<lambda>n. (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4254
    (foldl (\<lambda>r b. r * pochhammer b n) 1 bs * of_nat (fact n)))"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4255
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4256
lemma F_nth[simp]: "F as bs c $ n =
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4257
  (foldl (\<lambda>r a. r* pochhammer a n) 1 as * c^n) /
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4258
    (foldl (\<lambda>r b. r * pochhammer b n) 1 bs * of_nat (fact n))"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4259
  by (simp add: F_def)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4260
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4261
lemma foldl_mult_start:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4262
  fixes v :: "'a::comm_ring_1"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4263
  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
  4264
  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
  4265
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4266
lemma foldr_mult_foldl:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4267
  fixes v :: "'a::comm_ring_1"
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4268
  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
  4269
  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
  4270
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4271
lemma F_nth_alt:
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4272
  "F 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
  4273
    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
  4274
  by (simp add: foldl_mult_start foldr_mult_foldl)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4275
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4276
lemma F_E[simp]: "F [] [] c = E c"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4277
  by (simp add: fps_eq_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4278
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4279
lemma F_1_0[simp]: "F [1] [] c = 1/(1 - fps_const c * X)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4280
proof -
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4281
  let ?a = "(Abs_fps (\<lambda>n. 1)) oo (fps_const c * X)"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4282
  have th0: "(fps_const c * X) $ 0 = 0" by simp
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4283
  show ?thesis unfolding gp[OF th0, symmetric]
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4284
    by (auto simp add: fps_eq_iff pochhammer_fact[symmetric]
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  4285
      fps_compose_nth power_mult_distrib cond_value_iff setsum.delta' cong del: if_weak_cong)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4286
qed
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4287
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4288
lemma F_B[simp]: "F [-a] [] (- 1) = fps_binomial a"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4289
  by (simp add: fps_eq_iff gbinomial_pochhammer algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4290
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4291
lemma F_0[simp]: "F as bs c $ 0 = 1"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4292
  apply simp
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4293
  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
  4294
  apply auto
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4295
  apply (induct_tac as)
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4296
  apply auto
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4297
  done
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4298
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4299
lemma foldl_prod_prod:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4300
  "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
  4301
    foldl (\<lambda>r x. r * f x * g x) (v * w) as"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4302
  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
  4303
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4304
53196
942a1b48bb31 tuned proofs;
wenzelm
parents: 53195
diff changeset
  4305
lemma F_rec:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4306
  "F as bs c $ Suc n = ((foldl (\<lambda>r a. r* (a + of_nat n)) c as) /
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4307
    (foldl (\<lambda>r b. r * (b + of_nat n)) (of_nat (Suc n)) bs )) * F as bs c $ n"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4308
  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
  4309
  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
  4310
  unfolding foldl_prod_prod[unfolded foldl_mult_start] pochhammer_Suc
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4311
  apply (simp add: algebra_simps of_nat_mult)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4312
  done
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4313
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4314
lemma XD_nth[simp]: "XD a $ n = (if n = 0 then 0 else of_nat n * a$n)"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4315
  by (simp add: XD_def)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4316
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4317
lemma XD_0th[simp]: "XD a $ 0 = 0"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4318
  by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4319
lemma XD_Suc[simp]:" XD a $ Suc n = of_nat (Suc n) * a $ Suc n"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4320
  by simp
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4321
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4322
definition "XDp c a = XD a + fps_const c * a"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4323
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4324
lemma XDp_nth[simp]: "XDp c a $ n = (c + of_nat n) * a$n"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4325
  by (simp add: XDp_def algebra_simps)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4326
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  4327
lemma XDp_commute: "XDp b \<circ> XDp (c::'a::comm_ring_1) = XDp c \<circ> XDp b"
39302
d7728f65b353 renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents: 39198
diff changeset
  4328
  by (auto simp add: XDp_def fun_eq_iff fps_eq_iff algebra_simps)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4329
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4330
lemma XDp0 [simp]: "XDp 0 = 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
  4331
  by (simp add: fun_eq_iff fps_eq_iff)
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4332
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4333
lemma XDp_fps_integral [simp]: "XDp 0 (fps_integral a c) = X * a"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4334
  by (simp add: fps_eq_iff fps_integral_def)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4335
52891
b8dede3a4f1d tuned proofs;
wenzelm
parents: 51542
diff changeset
  4336
lemma F_minus_nat:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59862
diff changeset
  4337
  "F [- of_nat n] [- of_nat (n + m)] (c::'a::{field_char_0,field}) $ k =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4338
    (if k \<le> n then
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4339
      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
  4340
     else 0)"
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59862
diff changeset
  4341
  "F [- of_nat m] [- of_nat (m + n)] (c::'a::{field_char_0,field}) $ k =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4342
    (if k \<le> m then
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4343
      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
  4344
     else 0)"
32160
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4345
  by (auto simp add: pochhammer_eq_0_iff)
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4346
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4347
lemma setsum_eq_if: "setsum f {(n::nat) .. m} = (if m < n then 0 else f n + setsum f {n+1 .. m})"
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4348
  apply simp
57418
6ab1c7cb0b8d fact consolidation
haftmann
parents: 57129
diff changeset
  4349
  apply (subst setsum.insert[symmetric])
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4350
  apply (auto simp add: not_less setsum_head_Suc)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4351
  done
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4352
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4353
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
  4354
  by (cases n) (simp_all add: pochhammer_rec)
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4355
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  4356
lemma XDp_foldr_nth [simp]: "foldr (\<lambda>c r. XDp c \<circ> r) cs (\<lambda>c. XDp c a) c0 $ n =
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4357
    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
  4358
  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
  4359
63686057cbe8 Vandermonde vs Pochhammer; Hypergeometric series - very basic facts
chaieb
parents: 32157
diff changeset
  4360
lemma genric_XDp_foldr_nth:
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4361
  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
  4362
  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
  4363
    foldr (\<lambda>c r. (k c + of_nat n) * r) cs (g c0 a $ n)"
48757
1232760e208e tuned proofs;
wenzelm
parents: 47217
diff changeset
  4364
  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
  4365
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4366
lemma dist_less_imp_nth_equal:
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4367
  assumes "dist f g < inverse (2 ^ i)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4368
    and"j \<le> i"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4369
  shows "f $ j = g $ j"
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  4370
proof (rule ccontr)
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  4371
  assume "f $ j \<noteq> g $ j"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4372
  hence "f \<noteq> g" by auto
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4373
  with assms have "i < subdegree (f - g)"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4374
    by (simp add: split_if_asm dist_fps_def)
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 54230
diff changeset
  4375
  also have "\<dots> \<le> j"
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4376
    using \<open>f $ j \<noteq> g $ j\<close> by (intro subdegree_leI) simp_all
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  4377
  finally show False using \<open>j \<le> i\<close> by simp
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4378
qed
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4379
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4380
lemma nth_equal_imp_dist_less:
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4381
  assumes "\<And>j. j \<le> i \<Longrightarrow> f $ j = g $ j"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4382
  shows "dist f g < inverse (2 ^ i)"
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4383
proof (cases "f = g")
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4384
  case True
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4385
  then show ?thesis by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4386
next
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4387
  case False
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4388
  with assms have "dist f g = inverse (2 ^ subdegree (f - g))"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4389
    by (simp add: split_if_asm dist_fps_def)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4390
  moreover
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4391
  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
  4392
    by (intro subdegree_greaterI) simp_all
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4393
  ultimately show ?thesis by simp
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4394
qed
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4395
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4396
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
  4397
  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
  4398
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4399
instance fps :: (comm_ring_1) complete_space
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4400
proof
54681
8a8e6db7f391 tuned proofs;
wenzelm
parents: 54489
diff changeset
  4401
  fix X :: "nat \<Rightarrow> 'a fps"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4402
  assume "Cauchy X"
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4403
  obtain M where M: "\<forall>i. \<forall>m \<ge> M i. \<forall>j \<le> i. X (M i) $ j = X m $ j"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4404
  proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4405
    have "\<exists>M. \<forall>m \<ge> M. \<forall>j\<le>i. X M $ j = X m $ j" for i
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4406
    proof -
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4407
      have "0 < inverse ((2::real)^i)" by simp
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4408
      from metric_CauchyD[OF \<open>Cauchy X\<close> this] dist_less_imp_nth_equal
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4409
      show ?thesis by blast
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4410
    qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4411
    then show ?thesis using that by metis
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4412
  qed
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4413
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4414
  show "convergent X"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4415
  proof (rule convergentI)
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4416
    show "X ----> Abs_fps (\<lambda>i. X (M i) $ i)"
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4417
      unfolding tendsto_iff
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4418
    proof safe
61608
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4419
      fix e::real assume e: "0 < e"
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4420
      have "(\<lambda>n. inverse (2 ^ n) :: real) ----> 0" by (rule LIMSEQ_inverse_realpow_zero) simp_all
a0487caabb4a subdegree/shift/cutoff and Euclidean ring instance for formal power series
eberlm
parents: 61585
diff changeset
  4421
      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
  4422
        by (rule order_tendstoD)
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4423
      then obtain i where "inverse (2 ^ i) < e"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4424
        by (auto simp: eventually_sequentially)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4425
      have "eventually (\<lambda>x. M i \<le> x) sequentially"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4426
        by (auto simp: eventually_sequentially)
54452
f3090621446e tuned proofs;
wenzelm
parents: 54263
diff changeset
  4427
      then show "eventually (\<lambda>x. dist (X x) (Abs_fps (\<lambda>i. X (M i) $ i)) < e) sequentially"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4428
      proof eventually_elim
52902
7196e1ce1cd8 tuned proofs;
wenzelm
parents: 52891
diff changeset
  4429
        fix x
60501
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4430
        assume x: "M i \<le> x"
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4431
        have "X (M i) $ j = X (M j) $ j" if "j \<le> i" for j
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4432
          using M that by (metis nat_le_linear)
839169c70e92 tuned proofs;
wenzelm
parents: 60500
diff changeset
  4433
        with x have "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < inverse (2 ^ i)"
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4434
          using M by (force simp: dist_less_eq_nth_equal)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  4435
        also note \<open>inverse (2 ^ i) < e\<close>
51107
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4436
        finally show "dist (X x) (Abs_fps (\<lambda>j. X (M j) $ j)) < e" .
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4437
      qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4438
    qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4439
  qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4440
qed
3f9dbd2cc475 complete metric for formal power series
immler
parents: 49962
diff changeset
  4441
29911
c790a70a3d19 declare fps_nth as a typedef morphism; clean up instance proofs
huffman
parents: 29906
diff changeset
  4442
end